From e842ceaa03758b265341637df8775d3083f6a938 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Wed, 13 Mar 2024 01:15:51 +0100 Subject: [PATCH 01/20] Groupoid and Group Solvers --- Cubical/Data/List/Properties.agda | 5 + Cubical/Relation/Binary/Base.agda | 12 + Cubical/Tactics/GroupSolver/Example.agda | 36 + Cubical/Tactics/GroupSolver/Solver.agda | 44 + Cubical/Tactics/GroupoidSolver/Example.agda | 35 + .../Tactics/GroupoidSolver/Reflection.agda | 202 +++++ Cubical/Tactics/GroupoidSolver/Solver.agda | 754 ++++++++++++++++++ Cubical/WildCat/Base.agda | 53 ++ Cubical/WildCat/Functor.agda | 11 + 9 files changed, 1152 insertions(+) create mode 100644 Cubical/Tactics/GroupSolver/Example.agda create mode 100644 Cubical/Tactics/GroupSolver/Solver.agda create mode 100644 Cubical/Tactics/GroupoidSolver/Example.agda create mode 100644 Cubical/Tactics/GroupoidSolver/Reflection.agda create mode 100644 Cubical/Tactics/GroupoidSolver/Solver.agda diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index c44f9db1df..7fe83bfd03 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -179,3 +179,8 @@ length-map : ∀ {ℓA ℓB} {A : Type ℓA} {B : Type ℓB} → (f : A → B) → length (map f as) ≡ length as length-map f [] = refl length-map f (a ∷ as) = cong suc (length-map f as) + +lookupWithDefault : A → List A → ℕ → A +lookupWithDefault a [] _ = a +lookupWithDefault _ (x ∷ _) zero = x +lookupWithDefault a (x ∷ xs) (suc k) = lookupWithDefault a xs k diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index d583115a76..6fbfdca590 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -65,6 +65,9 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where isSym : Type (ℓ-max ℓ ℓ') isSym = (a b : A) → R a b → R b a + isSym' : Type (ℓ-max ℓ ℓ') + isSym' = {a b : A} → R a b → R b a + isAsym : Type (ℓ-max ℓ ℓ') isAsym = (a b : A) → R a b → ¬ R b a @@ -141,6 +144,15 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where symmetric : isSym transitive : isTrans + reflexive' : isRefl' + reflexive' = reflexive _ + + symmetric' : isSym' + symmetric' = symmetric _ _ + + transitive' : isTrans' + transitive' = transitive _ _ _ + isUniversalRel→isEquivRel : HeterogenousRelation.isUniversalRel R → isEquivRel isUniversalRel→isEquivRel u .isEquivRel.reflexive a = u a a isUniversalRel→isEquivRel u .isEquivRel.symmetric a b _ = u b a diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda new file mode 100644 index 0000000000..ac032f53fb --- /dev/null +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupSolver.Example where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure + +open import Cubical.Algebra.Group +open import Cubical.Tactics.GroupSolver.Solver + +private + variable + ℓ : Level + + +module example (G : Group ℓ) where + open GroupStr (snd G) + + module _ (p p' q q' r s : ⟨ G ⟩) where + + + pA pB pC : ⟨ G ⟩ + pA = (p · (1g · q)) · inv r + pB = (p · ((q · (inv ((q' · inv q') · + inv r · ((s · inv s) · r)) · inv q)) · (q · 1g))) · (inv r · 1g) + pC = (1g · p') · (((q · inv q) · (inv p' · p)) · (q · (inv q · (q · inv r)))) + + + pA=pB : pA ≡ pB + pA=pB = solveGroup G + + pA=pC : pA ≡ pC + pA=pC = solveGroup G + + pB=pC : pB ≡ pC + pB=pC = solveGroup G diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda new file mode 100644 index 0000000000..a2bdaaa024 --- /dev/null +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -0,0 +1,44 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupSolver.Solver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Unit +open import Cubical.Data.List as Li + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R + +open import Cubical.Tactics.Reflection +open import Cubical.WildCat.Base +open import Cubical.Tactics.GroupoidSolver.Solver + +open import Cubical.Algebra.Group + +private + variable + ℓ : Level + + +module _ (G : Group ℓ) where + open GroupStr (snd G) + Group→WildGroupoid : WildGroupoid ℓ-zero ℓ + WildCat.ob (WildGroupoid.wildCat Group→WildGroupoid) = Unit + WildCat.Hom[_,_] (WildGroupoid.wildCat Group→WildGroupoid) _ _ = ⟨ G ⟩ + WildCat.id (WildGroupoid.wildCat Group→WildGroupoid) = 1g + WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ + WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL + WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR + WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) + wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f + wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f + wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f + + +macro + solveGroup : R.Term → R.Term → R.TC Unit + solveGroup g-t = groupoidSolverMain true (R.def (quote Group→WildGroupoid ) (g-t v∷ [])) diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda new file mode 100644 index 0000000000..d0a19509cd --- /dev/null +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -0,0 +1,35 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupoidSolver.Example where + +open import Cubical.Foundations.Prelude + +open import Cubical.WildCat.Base +open import Cubical.Tactics.GroupoidSolver.Solver + +private + variable + ℓ ℓ' : Level + + +module example (WG : WildGroupoid ℓ ℓ') where + open WildGroupoid WG + + module _ (x y z w v : ob) (p p' : Hom[ x , y ]) (q q' : Hom[ y , z ]) + (r : Hom[ w , z ]) (s : Hom[ w , v ]) where + + + pA pB pC : Hom[ x , w ] + pA = (p ⋆ (id ⋆ q)) ⋆ inv r + pB = (p ⋆ ((q ⋆ (inv (inv r ⋆ ((s ⋆ inv s) ⋆ r)) ⋆ inv q)) ⋆ (q ⋆ id))) ⋆ (inv r ⋆ id) + pC = (id ⋆ p') ⋆ (((q ⋆ inv q) ⋆ (inv p' ⋆ p)) ⋆ (q ⋆ (inv q ⋆ (q ⋆ inv r)))) + + + pA=pB : pA ≡ pB + pA=pB = solveGroupoid WG + + pA=pC : pA ≡ pC + pA=pC = solveGroupoid WG + + pB=pC : pB ≡ pC + pB=pC = solveGroupoid WG diff --git a/Cubical/Tactics/GroupoidSolver/Reflection.agda b/Cubical/Tactics/GroupoidSolver/Reflection.agda new file mode 100644 index 0000000000..cabd7aaf38 --- /dev/null +++ b/Cubical/Tactics/GroupoidSolver/Reflection.agda @@ -0,0 +1,202 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupoidSolver.Reflection where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Agda.Builtin.Char +open import Agda.Builtin.String + +open import Cubical.Data.Bool +open import Cubical.Data.Nat +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Maybe + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.Reflection + +infixr 40 _<>_ + +_≟ℕ_ : ℕ → ℕ → Bool +x ≟ℕ x₁ = Dec→Bool (discreteℕ x x₁) + +_<>_ = primStringAppend + +_≟qn_ = R.primQNameEquality + +quotedMaybe→maybeTerm : R.Term → R.TC (Maybe (R.Term)) +quotedMaybe→maybeTerm (R.con (quote nothing) _) = R.returnTC nothing +quotedMaybe→maybeTerm (R.con (quote just) (_ ∷ _ ∷ varg x ∷ [])) = R.returnTC (just x) +quotedMaybe→maybeTerm t = R.typeError + (R.termErr t ∷ [(R.strErr "Not a Maybe!")]) + + +quotedList→ListOfTerms : R.Term → R.TC (List (R.Term)) +quotedList→ListOfTerms (R.con (quote []) _) = R.returnTC [] +quotedList→ListOfTerms (R.con (quote _∷_) (_ ∷ _ ∷ varg x ∷ varg xs ∷ [])) = + quotedList→ListOfTerms xs >>= (λ xs → R.returnTC (x ∷ xs)) +quotedList→ListOfTerms t = R.typeError + (R.termErr t ∷ [(R.strErr "Not a List!")]) + +blockIfContainsMeta : R.Term → R.TC Unit + +blockIfContainsMeta' : List (R.Arg R.Term) → R.TC Unit +blockIfContainsMeta' [] = R.returnTC _ +blockIfContainsMeta' (R.arg i x ∷ x₁) = + blockIfContainsMeta x >> blockIfContainsMeta' x₁ + + +blockIfContainsMeta (R.var _ args) = blockIfContainsMeta' args +blockIfContainsMeta (R.con _ args) = blockIfContainsMeta' args +blockIfContainsMeta (R.def _ args) = blockIfContainsMeta' args +blockIfContainsMeta (R.lam _ (R.abs s x)) = blockIfContainsMeta x +blockIfContainsMeta (R.pat-lam _ _) = R.typeError [(R.strErr "TODO : blockIfContainsMeta")] +blockIfContainsMeta (R.pi _ _) = R.typeError [(R.strErr "TODO : blockIfContainsMeta")] +blockIfContainsMeta (R.agda-sort _) = R.typeError [(R.strErr "TODO : blockIfContainsMeta")] +blockIfContainsMeta (R.lit l) = R.returnTC _ +blockIfContainsMeta (R.meta m _) = R.blockTC (R.blockerMeta m) +blockIfContainsMeta R.unknown = R.returnTC _ + +any? : List Bool → Bool +any? [] = false +any? (false ∷ x₁) = any? x₁ +any? (true ∷ x₁) = true + +containsMeta? : R.Term → Bool +containsMetaAny? : List (R.Arg R.Term) → Bool + +containsMeta? (R.var x args) = containsMetaAny? args +containsMeta? (R.con c args) = containsMetaAny? args +containsMeta? (R.def f args) = containsMetaAny? args +containsMeta? (R.lam v₁ (R.abs _ t)) = containsMeta? t +containsMeta? (R.pat-lam cs args) = containsMetaAny? args +containsMeta? (R.pi (R.arg _ a) (R.abs _ b)) = containsMeta? a or containsMeta? b +containsMeta? (R.agda-sort s) = false +containsMeta? (R.lit l) = false +containsMeta? (R.meta x x₁) = true +containsMeta? R.unknown = true + +containsMetaAny? [] = false +containsMetaAny? ((R.arg _ x) ∷ x₁) = containsMeta? x or containsMetaAny? x₁ + +catchOrSkip : ∀ {ℓ} {A : Type ℓ} → Bool → R.TC A → R.TC A → R.TC A +catchOrSkip true _ x = x +catchOrSkip false x y = R.catchTC x y + +uniqeAtoms : List R.Term → R.TC (List R.Term) +uniqeAtoms [] = R.returnTC [] +uniqeAtoms (x ∷ x₁) = do + notIn ← ensureNotIn x₁ + xs' ← uniqeAtoms x₁ + R.returnTC (if notIn then x ∷ xs' else xs') + where + ensureNotIn : List R.Term → R.TC Bool + ensureNotIn [] = R.returnTC true + ensureNotIn (x' ∷ xs) = + R.catchTC ( (R.unify x x' >> R.returnTC false)) + (ensureNotIn xs) + + +lookT : List R.Term → R.Term → R.TC ℕ +lookT [] _ = R.typeError [] +lookT (x ∷ x₂) x₁ = + R.catchTC ((R.unify x x₁ >> R.returnTC zero)) + (lookT x₂ x₁ >>= R.returnTC ∘ suc) + +quoteList : List R.Term → R.Term +quoteList [] = R.con (quote []) [] +quoteList (x ∷ x₁) = R.con (quote _∷_) + (varg x ∷ varg (quoteList x₁) ∷ []) + + +matchVarg : List (R.Arg R.Term) → R.TC R.Term +matchVarg (harg _ ∷ xs) = matchVarg xs +matchVarg (varg t ∷ []) = R.returnTC t +matchVarg _ = R.typeError [ R.strErr "matchV fail" ] + + +match2Vargs : List (R.Arg R.Term) → R.TC (R.Term × R.Term) +match2Vargs (harg _ ∷ xs) = match2Vargs xs +match2Vargs (varg t1 ∷ varg t2 ∷ []) = R.returnTC (t1 , t2) +match2Vargs _ = R.typeError [] + + +match3Vargs : List (R.Arg R.Term) → R.TC (R.Term × R.Term × R.Term) +match3Vargs (harg _ ∷ xs) = match3Vargs xs +match3Vargs (varg t1 ∷ varg t2 ∷ varg t3 ∷ []) = R.returnTC (t1 , t2 , t3) +match3Vargs _ = R.typeError [] + + + +inferEnds : R.Term → R.TC (R.Type × (R.Term × R.Term)) +inferEnds hole = do + ty ← R.inferType hole >>= wait-for-type + (eTy , (e0 , e1)) ← R.withNormalisation true $ pathTypeView ty + blockIfContainsMeta e0 + blockIfContainsMeta e1 + + R.returnTC (eTy , (e0 , e1)) + where + pathTypeView : R.Term → R.TC (R.Type × (R.Term × R.Term)) + pathTypeView (R.def (quote _≡_) l@(_ ∷ (harg ty) ∷ _)) = + do e ← match2Vargs l + R.returnTC (ty , e) + pathTypeView t = R.typeError (R.strErr "Not a Path Type! : " ∷ [ R.termErr t ]) + +digitsToSubscripts : Char → Char +digitsToSubscripts = λ where + '0' → '₀' ; '1' → '₁' ; '2' → '₂' ; '3' → '₃' ; '4' → '₄' ; '5' → '₅' + '6' → '₆' ; '7' → '₇' ; '8' → '₈' ; '9' → '₉' ; x → x + +mkNiceVar : ℕ → String +mkNiceVar k = "𝒙" <> + primStringFromList (map digitsToSubscripts $ primStringToList $ primShowNat k) + +record ToErrorPart {ℓ} (A : Type ℓ) : Type ℓ where + field + toErrorPart : A → R.ErrorPart + +open ToErrorPart + +infixr 5 _∷ₑ_ _∷nl_ +_∷ₑ_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart +_∷ₑ_ ⦃ tep ⦄ x = (toErrorPart tep x) ∷_ + + +instance + ToErrorPartString : ToErrorPart String + ToErrorPart.toErrorPart ToErrorPartString = R.strErr + + ToErrorPartTerm : ToErrorPart R.Term + ToErrorPart.toErrorPart ToErrorPartTerm = R.termErr + + ToErrorPartName : ToErrorPart R.Name + ToErrorPart.toErrorPart ToErrorPartName = R.nameErr + + +_∷nl_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart +_∷nl_ x y = x ∷ₑ "\n" ∷ₑ y + + +niceAtomList : List (R.Term) → List R.ErrorPart +niceAtomList = h 0 + where + h : _ → _ + h _ [] = [] + h k (x ∷ xs) = (mkNiceVar k <> " = ") ∷ₑ x ∷ₑ "\n" ∷ₑ h (suc k) xs + + +foldlM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (B → A → R.TC B) → B → List A → R.TC B +foldlM f b [] = R.returnTC b +foldlM f b (x ∷ xs) = f b x >>= (flip (foldlM f)) xs + +R∙ : R.Term → R.Term → R.Term +R∙ x y = R.def (quote _∙_) (x v∷ y v∷ [] ) + +Rrefl : R.Term +Rrefl = R.def (quote refl) [] diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda new file mode 100644 index 0000000000..2c16c3ef51 --- /dev/null +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -0,0 +1,754 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupoidSolver.Solver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function as Fu +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.GroupoidSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + +-- open import Cubical.WildCat.WGE +open import Cubical.WildCat.Base + + +private + variable + ℓ ℓ' : Level + +unfoldMaybe : ℕ → ∀ {ℓ} {A : Type ℓ} → (A → Maybe A) → A → List A +unfoldMaybe zero _ _ = [] +unfoldMaybe (suc x) f a = + Mb.rec [] (λ a' → a' ∷ unfoldMaybe x f a') (f a) + +data 𝑵𝒐𝒅𝒆 : Type where + 𝒐𝒑 𝒂𝒕𝒐𝒎 𝒄𝒐𝒏𝒔 𝒊𝒏𝒗 𝒊𝒏𝒗𝒐𝒍 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 : 𝑵𝒐𝒅𝒆 + +module Nodes (ob : Type ℓ) (Hom[_,_] : ob → ob → Type ℓ') (hasInvs : Bool) where + + 𝒏 : ℕ → 𝑵𝒐𝒅𝒆 → Bool + + 𝒏 0 𝒐𝒑 = true + 𝒏 0 𝒊𝒏𝒗 = hasInvs + + 𝒏 0 𝒄𝒐𝒏𝒔 = false + 𝒏 0 𝒂𝒕𝒐𝒎 = true + 𝒏 0 𝒊𝒏𝒗𝒐𝒍 = false + 𝒏 0 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 = false + + 𝒏 (suc _) 𝒊𝒏𝒗 = false + 𝒏 1 𝒐𝒑 = true + 𝒏 1 𝒄𝒐𝒏𝒔 = false + 𝒏 1 𝒂𝒕𝒐𝒎 = true + 𝒏 1 𝒊𝒏𝒗𝒐𝒍 = false + 𝒏 (suc _) 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 = hasInvs + + 𝒏 (suc (suc _)) 𝒐𝒑 = false + 𝒏 (suc (suc _)) 𝒄𝒐𝒏𝒔 = true + 𝒏 (suc (suc _)) 𝒊𝒏𝒗𝒐𝒍 = hasInvs + 𝒏 (suc (suc _)) 𝒂𝒕𝒐𝒎 = false + + + open BinaryRelation Hom[_,_] public + + + module _ (k : ℕ) where + N : 𝑵𝒐𝒅𝒆 → Type + N = Bool→Type Fu.∘ 𝒏 k + + data Atom : ob → ob → Type (ℓ-max ℓ ℓ') where + a⟦_⟧ : ∀ {x y} → Hom[ x , y ] → Atom x y + a⟦_⟧⁻ : ∀ {x y} → {invG : N 𝒊𝒏𝒗𝑨𝒕𝒐𝒎} → Hom[ y , x ] → {Bool→Type hasInvs} → Atom x y + + Atom→𝟚×H : ∀ {x y} → Atom x y → Σ Bool λ { true → Hom[ x , y ] ; false → Hom[ y , x ] } + Atom→𝟚×H a⟦ x ⟧ = true , x + Atom→𝟚×H a⟦ x ⟧⁻ = false , x + + data Node' : ob → ob → Type (ℓ-max ℓ ℓ') + data Node : ob → ob → Type (ℓ-max ℓ ℓ') where + idN : ∀ {x} → Node x x + atomN : ∀ {x y} → {aGuard : N 𝒂𝒕𝒐𝒎} → Atom x y → Node x y + no : ∀ {x y} → Node' x y → Node x y + + data Node' where + _∷N_ : ∀ {x y z} → {∷Guard : N 𝒄𝒐𝒏𝒔} → Node x y → Atom y z → Node' x z + _⋆N_ : ∀ {x y z} → {⋆Guard : N 𝒐𝒑} → Node x y → Node y z → Node' x z + invN : ∀ {x y} → {invGuard : N 𝒊𝒏𝒗} → Node y x → {Bool→Type hasInvs} → Node' x y + involN : ∀ {x y z} → {g : N 𝒊𝒏𝒗𝒐𝒍} → Node x y → Atom y z → {Bool→Type hasInvs} → Node' x y + + + _a⤋_ : ∀ k → ∀ {x y} → Atom k x y → Atom (suc k) x y + k a⤋ a⟦ x ⟧ = a⟦ x ⟧ + k a⤋ (a⟦ x ⟧⁻ {g}) = a⟦_⟧⁻ {_} {_} {_} {g} x {g} + + + _N2++_ : ∀ {x y z k} → Node (suc (suc k)) x y + → Node (suc (suc k)) y z + → Node (suc (suc k)) x z + x N2++ idN = x + x N2++ atomN x₁ = no (x ∷N x₁) + x N2++ no (x₁ ∷N x₂) = no ((x N2++ x₁) ∷N x₂) + x N2++ no (involN {g = g} x₁ x₂ {uuu}) = no (involN {g = g} (x N2++ x₁) x₂ {uuu}) + + + invAtom : ∀ {x y} k {k'} → (Bool→Type hasInvs) → Atom k x y → Atom (suc k') y x + invAtom k x a⟦ x₁ ⟧ = a⟦_⟧⁻ {_} {_} {_} {x} x₁ {x} + invAtom k x a⟦ x₁ ⟧⁻ = a⟦ x₁ ⟧ + + invNode : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node k x y → Node k y x + invNode' : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node' k x y → Node k y x + + invNode' {k = k} {x} (_⋆N_ {⋆Guard = ⋆Guard} x₁ x₂) = + no (_⋆N_ {⋆Guard = ⋆Guard} (invNode {hi = x} x₂) (invNode {hi = x} x₁)) + invNode' {k = k} {x} (invN {invGuard = invGuard} x₁) = x₁ + invNode' {k = suc (suc k)} {x} (involN {g = gg} x₁ x₂ {g}) = + no (involN {g = gg} idN (x₂) {g}) N2++ invNode {hi = x} x₁ + invNode' {k = suc (suc k)} {x} (x₁ ∷N x₂) = + no (idN ∷N invAtom _ x x₂) N2++ invNode {hi = x} x₁ + invNode {k = k} {x} idN = idN + invNode {k = zero} {x} (atomN {aGuard = g} x₁) = + no (invN {invGuard = x} (atomN {aGuard = g} x₁) {x}) + invNode {k = suc k} {x} (atomN {aGuard = g} x₁) = + atomN {aGuard = g} (invAtom _ x x₁) + invNode {k = k} {x} (no x₁) = invNode' {hi = x} x₁ + + + invNode* : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node k x y → Node k y x + invNode* {k = zero} {hi} f = no (invN {invGuard = hi} f {hi}) + invNode* {k = suc k} {hi} = invNode {k = suc k} {hi} + + + len : ∀ {x y} → Node 2 x y → ℕ + len idN = 0 + len (no (x ∷N x₁)) = suc (len x) + len (no (involN x x₁)) = suc (suc (len x)) + + module _ {k} (showH : {a₀ a₁ : ob} → (p : Hom[ a₀ , a₁ ]) → String) where + + showA showA' : {a₀ a₁ : ob} → Atom k a₀ a₁ → String + showA a⟦ x ⟧ = showH x + showA a⟦ x ⟧⁻ = "(" <> showH x <> ")⁻¹" + showA' a⟦ x ⟧ = "(" <> showH x <> ")⁻¹" + showA' a⟦ x ⟧⁻ = showH x + + showN : {a₀ a₁ : ob} → Node k a₀ a₁ → String + showN' : {a₀ a₁ : ob} → Node' k a₀ a₁ → String + + showN idN = "id" + showN (atomN x) = showA x + showN (no x) = showN' x + + showN' (x ∷N x₁) = showN x <> "⋆" <> showA x₁ + showN' (x ⋆N x₁) = "(" <> showN x <> "⋆" <> showN x₁ <> ")" + showN' (invN x) = "(" <> showN x <> ")⁻¹" + showN' (involN x x₁ {v}) = showN x <> + "⋆⟦" <> showA x₁ <> "⋆" <> showA' x₁ <> "⟧" + + + _⤋'_ : ∀ k → ∀ {x y} → Node' k x y → Node (suc k) x y + + _⤋_ : ∀ k → ∀ {x y} → Node k x y → Node (suc k) x y + k ⤋ idN = idN + zero ⤋ atomN x = atomN (0 a⤋ x) + suc zero ⤋ atomN x = no (idN ∷N (_ a⤋ x)) + k ⤋ no x = (k ⤋' x) + + suc (suc k) ⤋' (x ∷N x₁) = no ((suc (suc k) ⤋ x) ∷N (suc (suc k) a⤋ x₁)) + zero ⤋' (x ⋆N x₁) = no ((0 ⤋ x) ⋆N (0 ⤋ x₁)) + suc zero ⤋' (x ⋆N x₁) = + (suc zero ⤋ x) N2++ (suc zero ⤋ x₁) + zero ⤋' invN x {hi} = invNode {hi = hi} (zero ⤋ x) + + suc (suc k) ⤋' involN {_} {_} {_} {zz} x x₁ {zzz} = + no (involN {_} {_} {_} {_} {zz} (suc (suc k) ⤋ x) (suc (suc k) a⤋ x₁) {zzz}) + + _⤋⁺_ : ∀ {k} m → ∀ {x y} → Node k x y → Node (m + k) x y + zero ⤋⁺ x = x + suc m ⤋⁺ x = (m + _) ⤋ (m ⤋⁺ x) + + red[_,_] : ∀ k → ∀ {x y} → Node k x y → Node k x y + red[ zero , f ] = f + red[ suc zero , f ] = f + red[ suc (suc k) , idN ] = idN + red[ suc (suc k) , no (x ∷N x₁) ] = + no (red[ suc (suc k) , x ] ∷N x₁) + red[ suc (suc k) , no (involN x x₁) ] = red[ suc (suc k) , x ] + + + + module Ev (id : isRefl') + (_⋆_ : isTrans') + (inv : {_ : Bool→Type hasInvs} → isSym') where + + + eva[_] : ∀ {k} → ∀ {x y} → Atom k x y → Hom[ x , y ] + eva[ a⟦ x ⟧ ] = x + eva[ a⟦ x ⟧⁻ {invGuard} ] = inv {invGuard} x + + ev[_] : ∀ {k} → ∀ {x y} → Node k x y → Hom[ x , y ] + ev[_]b : ∀ {k} → ∀ {x y} → Node k x y → Bool → Hom[ x , y ] + + ev[ f ] = ev[ f ]b false + + + ev[ idN ]b _ = id + ev[ atomN x ]b _ = eva[ x ] + ev[ no (xs ∷N x) ]b b = ev[ xs ]b b ⋆ eva[ x ] + ev[ no (x ⋆N x₁) ]b b = ev[ x ]b b ⋆ ev[ x₁ ]b b + ev[ no (invN x {invGuard}) ]b b = inv {invGuard} (ev[ x ]b b) + ev[ no (involN x x₁ {invGuard}) ]b false = (ev[ x ] ⋆ eva[ x₁ ]) ⋆ + eva[ invAtom _ {1} invGuard x₁ ] + ev[ no (involN x x₁) ]b true = ev[ x ]b true + + module Ev' (⋆Assoc : ∀ {u v w x} + (f : Hom[ u , v ]) + (g : Hom[ v , w ]) + (h : Hom[ w , x ]) + → (f ⋆ g) ⋆ h ≡ f ⋆ (g ⋆ h)) + (⋆IdR : ∀ {x y} (f : Hom[ x , y ]) → f ⋆ id ≡ f) + (⋆IdL : ∀ {x y} (f : Hom[ x , y ]) → id ⋆ f ≡ f) + (⋆InvR : ∀ {x y hi} → (f : Hom[ x , y ]) → f ⋆ inv {hi} f ≡ id) + (⋆InvL : ∀ {x y hi} → (f : Hom[ x , y ]) → inv {hi} f ⋆ f ≡ id) + where + + evInv : ∀ {k} → ∀ {x y hi} → (f : Atom k x y) → + id ≡ (eva[ f ] ⋆ eva[ invAtom k {1} hi f ]) + evInv a⟦ x ⟧ = sym (⋆InvR x) + evInv a⟦ x ⟧⁻ = sym (⋆InvL x) + + ev[_]≡ : ∀ {k} → ∀ {x y} → (f : Node k x y) → ev[ f ]b false ≡ ev[ f ]b true + ev[ idN ]≡ = refl + ev[ atomN x ]≡ = refl + ev[ no (x ∷N x₁) ]≡ = cong (_⋆ eva[ x₁ ]) ev[ x ]≡ + ev[ no (x ⋆N x₁) ]≡ = cong₂ _⋆_ ev[ x ]≡ ev[ x₁ ]≡ + ev[ no (invN x) ]≡ = cong inv ev[ x ]≡ + ev[ no (involN x x₁) ]≡ = + ⋆Assoc _ _ _ + ∙∙ cong₂ _⋆_ ev[ x ]≡ (sym (evInv x₁)) + ∙∙ ⋆IdR _ + + ev[_]₂≡ = ev[_]≡ {k = 2} + + + evN2++b : ∀ {k} {a₀ a₁ a₂ : ob} b → (x : Node (suc (suc k)) a₀ a₁) + → (x₁ : Node (suc (suc k)) a₁ a₂) + → ev[ x ]b b ⋆ ev[ x₁ ]b b ≡ ev[ x N2++ x₁ ]b b + evN2++b b x idN = ⋆IdR _ + evN2++b b x (no (x₁ ∷N x₂)) = + sym (⋆Assoc _ _ _) + ∙ cong (_⋆ eva[ x₂ ]) (evN2++b b x x₁) + evN2++b false x (no (involN x₁ x₂)) = + sym (⋆Assoc _ _ _) + ∙∙ congS (_⋆ eva[ invAtom _ _ x₂ ]) (sym (⋆Assoc _ _ _)) + ∙∙ congS (λ y → ((y ⋆ eva[ x₂ ]) ⋆ eva[ invAtom _ _ x₂ ])) (evN2++b false x x₁) + + evN2++b true x (no (involN x₁ x₂)) = evN2++b true x x₁ + + evN2++ : ∀ {k} {a₀ a₁ a₂ : ob} → (x : Node (suc (suc k)) a₀ a₁) + → (x₁ : Node (suc (suc k)) a₁ a₂) + → ev[ x ] ⋆ ev[ x₁ ] ≡ ev[ x N2++ x₁ ] + evN2++ = evN2++b false + + module Ev'' (id≡inv-id : ∀ {x hi} → id {x} ≡ inv {hi} id) + (involInv : ∀ {x y hi hi'} → (f : Hom[ x , y ]) → + inv {hi} (inv {hi'} f) ≡ f) + (distInv : ∀ {x y z hi} + (f : Hom[ x , y ]) + (g : Hom[ y , z ]) + → inv {hi} (f ⋆ g) ≡ inv {hi} g ⋆ inv {hi} f) + where + + inv-eva : ∀ {x y} {k} {k'} a {hi} → + inv {hi} {x} {y} (eva[_] {k} a) ≡ eva[ invAtom (k) {k'} hi a ] + inv-eva a⟦ x ⟧ = refl + inv-eva a⟦ x ⟧⁻ {hi} = involInv x + + invAtomInvol : ∀ {x y} k {hi hi'} → (x : Atom (suc (suc k)) x y) → + (eva[ invAtom 2 {suc k} (hi) (invAtom (suc (suc k)) {1} hi' x) ]) ≡ (eva[ x ]) + invAtomInvol k a⟦ x ⟧ = refl + invAtomInvol k {hi} (a⟦ x ⟧⁻ {hi'}) i = inv {isPropBool→Type hi hi' i} x + + + ev[_]≡inv : ∀ {k} → ∀ {x y} {hi} b → (f : Node k x y) → + inv {hi} (ev[ f ]b b) ≡ ev[ invNode* {hi = hi} f ]b b + ev[_]≡inv {zero} b f = refl + ev[_]≡inv {suc k} {hi = hi} b idN = sym (id≡inv-id {hi = hi}) + ev[_]≡inv {suc zero} {hi = hi} b (atomN x) = inv-eva x {hi} + ev[_]≡inv {suc zero} b (no (x ⋆N x₁)) = + distInv (ev[ x ]b b) (ev[ x₁ ]b b) ∙ + cong₂ _⋆_ (ev[_]≡inv b x₁) (ev[_]≡inv b x) + ev[_]≡inv {suc (suc k)} {hi = hi} b (no (x ∷N x₁)) = + distInv (ev[ x ]b b) (eva[ x₁ ]) + ∙∙ cong₂ _⋆_ (λ i → ⋆IdL (inv-eva {k' = suc k} x₁ {hi} i) (~ i)) + (ev[_]≡inv b x) + ∙∙ evN2++b b (no (idN ∷N invAtom (suc (suc k)) hi x₁)) (invNode x) + + ev[_]≡inv {suc (suc k)} {hi = hi} false (no (involN x x₁ {hi'})) = + distInv (ev[ x ] ⋆ eva[ x₁ ]) (eva[ invAtom (suc (suc k)) hi' x₁ ]) + ∙∙ + cong₂ _⋆_ + (inv-eva {k' = suc k} (invAtom (suc (suc k)) hi' x₁)) + (distInv ev[ x ] eva[ x₁ ]) ∙ + (λ i → ⋆Assoc (⋆IdL (invAtomInvol k {hi} {hi'} x₁ i) (~ i)) + (inv-eva {k' = 1} x₁ {isPropBool→Type hi hi' i} i) + (ev[_]≡inv {suc (suc k)} {hi = hi} false x i) (~ i) ) + ∙∙ evN2++b false (no (involN idN x₁)) (invNode x) + + ev[_]≡inv {suc (suc k)} true (no (involN x x₁)) = + ev[_]≡inv true x + ∙∙ sym (⋆IdL _) + ∙∙ evN2++b true (no (involN idN x₁)) (invNode x) + + + + + data NodeCase : {a₀ a₁ : ob} → Hom[ a₀ , a₁ ] → Type (ℓ-max ℓ ℓ') where + idCase : ∀ {x} → NodeCase (id {a = x}) + opCase : ∀ {x y z : _} → (p : Hom[ x , y ]) (q : Hom[ y , z ]) → NodeCase (p ⋆ q) + invCase : ∀ {x y : _} → {hi : Bool→Type hasInvs} (p : Hom[ x , y ]) → NodeCase (inv {hi} p) + + +module _ (WC : WildCat ℓ ℓ') where + + open WildCat WC + + module WGi hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) {hi : Bool→Type hasInvs} where + open WildGroupoid (record { wildCat = WC ; isWildGroupoid = iwg hi }) public + using (inv ; ⋆InvR; ⋆InvL; id≡inv-id; distInv; invol-inv) + + invol-inv' : ∀ hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) + {hi hi' : Bool→Type hasInvs} {x y : WC .WildCat.ob} + (f : Hom[ x , y ]) → + WildGroupoid.inv + (record { wildCat = WC ; isWildGroupoid = iwg hi }) + (WildGroupoid.inv + (record { wildCat = WC ; isWildGroupoid = iwg hi' }) f) + ≡ f + invol-inv' true iwg {hi} {hi'} = WGi.invol-inv true iwg + + module WCTerm hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) {hi : Bool→Type hasInvs} where + + open WGi hasInvs iwg + + open Nodes ob Hom[_,_] hasInvs public + open Ev id _⋆_ (λ {ig} → wildIsIso.inv' Fu.∘ iwg ig) public + + + module WEv' = Ev' ⋆Assoc ⋆IdR ⋆IdL ⋆InvR ⋆InvL + + + module WEv'' = WEv'.Ev'' (sym (id≡inv-id)) (invol-inv' hasInvs iwg) distInv + + + eva⤋ : ∀ k {a₀ a₁ : ob} → ∀ (h : Atom k a₀ a₁) → eva[ h ] ≡ eva[ k a⤋ h ] + eva⤋ k a⟦ x ⟧ = refl + eva⤋ k a⟦ x ⟧⁻ = refl + + + invAtom⤋ : ∀ k k' {hi} {a₀ a₁ : ob} → ∀ (h : Atom (suc (suc k)) a₀ a₁) → + eva[ invAtom (suc (suc k)) {k'} hi h ] + ≡ eva[ invAtom (suc (suc (suc k))) {k'} hi (suc (suc k) a⤋ h) ] + invAtom⤋ k k' Nodes.a⟦ x ⟧ = refl + invAtom⤋ k k' Nodes.a⟦ x ⟧⁻ = refl + + + ev⤋' : ∀ k {a₀ a₁ : ob} → ∀ (f' : Node' k a₀ a₁) → ev[ no f' ] ≡ ev[ k ⤋' f' ] + ev⤋ : ∀ k {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ k ⤋ f ] + + + + ev⤋ k idN = refl + ev⤋ zero (Nodes.atomN Nodes.a⟦ x ⟧) = refl + ev⤋ (suc zero) (Nodes.atomN Nodes.a⟦ x ⟧) = sym (⋆IdL x) + ev⤋ (suc zero) (Nodes.atomN (Nodes.a⟦ x ⟧⁻)) = sym ((⋆IdL _)) + ev⤋ k (no x) = ev⤋' k x + + ev⤋' (suc (suc k)) (x ∷N x₁) = + cong₂ _⋆_ (ev⤋ (suc (suc k)) x) (eva⤋ (suc (suc k)) x₁ ) + ev⤋' zero (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ zero x) (ev⤋ zero x₁) + ev⤋' (suc zero) (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ 1 x) (ev⤋ 1 x₁) ∙ + WEv'.evN2++ (1 ⤋ x) (1 ⤋ x₁) + ev⤋' zero (invN x {hi}) = + cong inv (ev⤋ zero x) ∙ WEv''.ev[_]≡inv false (0 ⤋ x) -- enInv1Node x hi + ev⤋' (suc (suc k)) (Nodes.involN x x₁ {hi}) = + cong₂ _⋆_ (λ i → (ev⤋ (suc (suc k)) x i ⋆ (eva⤋ (suc (suc k)) x₁) i)) + (invAtom⤋ k 1 {hi} x₁) + + ev⤋⁺ : ∀ {k} m {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ m ⤋⁺ f ] + ev⤋⁺ zero f = refl + ev⤋⁺ (suc m) f = ev⤋⁺ m f ∙ ev⤋ (m + _) (m ⤋⁺ f) + + ev⤋² = ev⤋⁺ {0} 2 + + + +module WGTerm (WG : WildGroupoid ℓ ℓ') where + open WCTerm (WildGroupoid.wildCat WG) true (λ _ → WildGroupoid.isWildGroupoid WG) public + + open WGi (WildGroupoid.wildCat WG) true (λ _ → WildGroupoid.isWildGroupoid WG) + + open WildGroupoid WG hiding (⋆InvR ; ⋆InvL) + open Ev' ⋆Assoc ⋆IdR ⋆IdL ⋆InvR ⋆InvL public + +module _ (A : Type ℓ) where + module Expr = Nodes Unit (λ _ _ → A) + + module DecNodes (_≟A_ : Discrete A) where + + AtomTo𝟚×A : Expr.Atom true 2 _ _ → (Bool × A) + AtomTo𝟚×A Nodes.a⟦ x ⟧ = true , x + AtomTo𝟚×A Nodes.a⟦ x ⟧⁻ = false , x + + + mbRed : Expr.Node true 2 _ _ → Maybe (Expr.Node true 2 _ _) + mbRed Nodes.idN = nothing + mbRed (Nodes.no (Nodes.idN Nodes.∷N x₁)) = nothing + mbRed (Nodes.no (x'@(Nodes.no (x Nodes.∷N x₂)) Nodes.∷N x₁)) = + decRec (λ _ → just $ Nodes.no (Nodes.involN (Mb.rec x (idfun _) (mbRed x)) x₂) ) + (λ _ → map-Maybe (Nodes.no ∘ Nodes._∷N x₁) (mbRed x')) + (discreteΣ 𝟚._≟_ (λ _ → _≟A_) (AtomTo𝟚×A x₂) (AtomTo𝟚×A (Expr.invAtom true 2 _ x₁))) + mbRed (Nodes.no (Nodes.no (Nodes.involN x x₂) Nodes.∷N x₁)) = nothing + mbRed (Nodes.no (Nodes.involN x x₁)) = nothing + + redList : Expr.Node true 2 _ _ → List (Expr.Node true 2 _ _) + redList x = unfoldMaybe (Expr.len true x) (mbRed ∘ Expr.red[_,_] true 2) x + +redListℕ = DecNodes.redList ℕ discreteℕ + +mapExpr : ∀ {A : Type ℓ} {A' : Type ℓ'} {b k} + → (A → A') + → Expr.Node A b k _ _ + → Expr.Node A' b k _ _ +mapExpr {A = A} {A'} {b} {k} f = w + where + wa : Expr.Atom A b k _ _ → Expr.Atom A' b k _ _ + wa Nodes.a⟦ x ⟧ = Nodes.a⟦ f x ⟧ + wa (Nodes.a⟦_⟧⁻ {invG = ig} x {g}) = Nodes.a⟦_⟧⁻ {invG = ig} (f x) {g} + + w : Expr.Node A b k _ _ → Expr.Node A' b k _ _ + w Nodes.idN = Nodes.idN + w (Nodes.atomN {aGuard = ag} x) = Nodes.atomN {aGuard = ag} (wa x) + w (Nodes.no (Nodes._∷N_ {∷Guard = g} x x₁)) = + Nodes.no (Nodes._∷N_ {∷Guard = g} (w x) (wa x₁)) + w (Nodes.no (Nodes._⋆N_ {⋆Guard = g} x x₁)) = + Nodes.no (Nodes._⋆N_ {⋆Guard = g} (w x) (w x₁)) + w (Nodes.no (Nodes.invN {invGuard = invGuard} x {g})) = + (Nodes.no (Nodes.invN {invGuard = invGuard} (w x) {g})) + w (Nodes.no (Nodes.involN {g = g} x x₁ {g'})) = + Nodes.no (Nodes.involN {g = g} (w x) (wa x₁) {g'}) + + +mapExprQ : ∀ {A : Type ℓ} {b k} + → (A → R.Term) → Expr.Node A b k _ _ → R.Term +mapExprQ {A = A} {b} {k} f = w + where + wa : Expr.Atom A b k _ _ → R.Term + wa Nodes.a⟦ x ⟧ = R.con (quote Nodes.a⟦_⟧) (f x v∷ []) + wa Nodes.a⟦ x ⟧⁻ = R.con (quote Nodes.a⟦_⟧⁻) (f x v∷ []) + + w : Expr.Node A b k _ _ → R.Term + w' : Expr.Node' A b k _ _ → R.Term + w' (x Nodes.∷N x₁) = + R.con (quote Nodes._∷N_) (w x v∷ wa x₁ v∷ []) + w' (x Nodes.⋆N x₁) = + R.con (quote Nodes._⋆N_) (w x v∷ w x₁ v∷ []) + w' (Nodes.invN x) = + R.con (quote Nodes.invN) (w x v∷ []) + w' (Nodes.involN x x₁) = + R.con (quote Nodes.involN) (w x v∷ wa x₁ v∷ []) + + w Nodes.idN = R.con (quote Nodes.idN) [] + w (Nodes.atomN x) = R.con (quote Nodes.atomN) (wa x v∷ []) + w (Nodes.no x) = R.con (quote Nodes.no) (w' x v∷ []) + + +ExprAccumM : ∀ {A : Type ℓ} {A' : Type ℓ'} {ℓs} {S : Type ℓs} {b k} + → (S → A → R.TC (S × A')) → S + → Expr.Node A b k _ _ + → R.TC (S × Expr.Node A' b k _ _) +ExprAccumM {A = A} {A'} {S = S} {b} {k} f = w + where + open Expr + + wa : S → Expr.Atom A b k _ _ → R.TC (S × Expr.Atom A' b k _ _) + wa s a⟦ x ⟧ = (λ (s' , x') → s' , a⟦ x' ⟧) <$> f s x + wa s (a⟦_⟧⁻ {invG = g'} x {g}) = (λ (s' , x') → s' , (a⟦_⟧⁻ {invG = g'} x' {g})) <$> f s x + + w : S → Expr.Node A b k _ _ → R.TC (S × Expr.Node A' b k _ _) + w' : S → Expr.Node' A b k _ _ → R.TC (S × Expr.Node A' b k _ _) + w s idN = R.returnTC (s , idN) + w s (atomN {aGuard = ag} a) = + (λ (s' , a') → s' , atomN {aGuard = ag} a') <$> wa s a + w s (no x) = w' s x + + w' s (_∷N_ {∷Guard = g} x x₁) = do + (s' , x') ← w s x + (s' , x₁') ← wa s' x₁ + R.returnTC (s' , no (_∷N_ {∷Guard = g} x' x₁')) + + w' s (_⋆N_ {⋆Guard = g} x x₁) = do + (s' , x') ← w s x + (s' , x₁') ← w s' x₁ + R.returnTC (s' , no (_⋆N_ {⋆Guard = g} x' x₁')) + + w' s (invN {invGuard = g'} x {g}) = + (λ (s' , x') → s' , (no (invN {invGuard = g'} x' {g}))) <$> w s x + w' s (involN x x₁) = w s x + + +opCase' : ∀ (WG : WildGroupoid ℓ ℓ') {x y z} f g → + WGTerm.NodeCase WG {a₀ = x} {z} _ +opCase' WG {x} {y} {z} f g = WGTerm.opCase {WG = WG} {y = y} f g + + +invCase' : ∀ (WG : WildGroupoid ℓ ℓ') {x y} f → + WGTerm.NodeCase WG {a₀ = y} {x} _ +invCase' WG {x} {y} f = WGTerm.invCase {WG = WG} {x = x} {y} f + +id' : (WG : WildGroupoid ℓ ℓ') → ∀ {x} → WildGroupoid.Hom[_,_] WG x x +id' WG = WildGroupoid.id WG + +inv' : (WG : WildGroupoid ℓ ℓ') → ∀ {x y} → WildGroupoid.Hom[_,_] WG x y → WildGroupoid.Hom[_,_] WG y x +inv' WG = WildGroupoid.inv WG + + +⋆' : (WG : WildGroupoid ℓ ℓ') → ∀ {x y z} → WildGroupoid.Hom[_,_] WG x y → WildGroupoid.Hom[_,_] WG y z → WildGroupoid.Hom[_,_] WG x z +⋆' WG = WildGroupoid._⋆_ WG + + +module ETerm = Expr R.Term + +module _ (WGterm : R.Term) where + module EvTerm = ETerm.Ev true + (R.def (quote id') (WGterm v∷ [])) + (λ x y → (R.def (quote ⋆') (WGterm v∷ x v∷ y v∷ []))) + (λ x → (R.def (quote inv') (WGterm v∷ x v∷ []))) + +module Eℕ = Expr ℕ true + +NodeTerm : Bool → ℕ → Type ℓ-zero +NodeTerm = λ b k → Expr.Node R.Term b k tt tt + + + + +module tryGE (tG : R.Term) where + + tryG : ℕ → R.Term → R.TC (NodeTerm true 0) + + try1g : R.Term → R.TC (NodeTerm true 0) + try1g t = do + _ ← R.unify t (R.def (quote id') [ varg tG ]) + R.returnTC (ETerm.idN) + + tryOp : ℕ → R.Term → R.TC (NodeTerm true 0) + tryOp zero _ = R.typeError [] + tryOp (suc k) t = do + tm ← R.withNormalisation true $ R.checkType (R.def (quote opCase') + (varg tG ∷ varg R.unknown ∷ [ varg R.unknown ])) + (R.def (quote WGTerm.NodeCase) ((varg tG) ∷ [ varg t ])) + (t1 , t2) ← h tm + t1' ← tryG k t1 + t2' ← tryG k t2 + R.returnTC (ETerm.no (t1' ETerm.⋆N t2')) + + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con _ l) = match2Vargs l + h t = R.typeError [] + + tryInv : ℕ → R.Term → R.TC (NodeTerm true 0) + tryInv zero _ = R.typeError [] + tryInv (suc k) t = do + tm ← R.withNormalisation true $ + (R.checkType (R.def (quote invCase') + ((varg tG) ∷ [ varg R.unknown ])) (R.def (quote WGTerm.NodeCase) + ((varg tG) ∷ [ varg t ]))) + R.debugPrint "tryInv" 30 ([ R.strErr "\n ---- \n" ]) + R.debugPrint "tryInv" 30 ([ R.termErr t ]) + R.debugPrint "tryInv" 30 ([ R.termErr tm ]) + t1 ← h tm + t1' ← tryG k t1 + R.returnTC (ETerm.no (ETerm.invN t1')) + + where + h' : List (R.Arg R.Term) → R.TC (R.Term) + h' (harg _ ∷ xs) = h' xs + h' (varg t1 ∷ []) = R.returnTC t1 + h' _ = do R.debugPrint "aiu" 3 ([ R.strErr "\n vvv \n" ]) + R.debugPrint "aiu" 33 ([ R.termErr t ]) + R.typeError [ R.strErr "yyy" ] + + h : R.Term → R.TC (R.Term) + h (R.con _ l) = h' l + h t = do R.debugPrint "aiu" 33 ([ R.strErr "\n uuu \n" ]) + R.debugPrint "aiu" 33 ([ R.termErr t ]) + R.typeError [ R.strErr "xxxx" ] + + + atom : R.Term → R.TC (NodeTerm true 0) + atom x = R.returnTC $ ETerm.atomN ETerm.a⟦ x ⟧ + + + tryG zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryG (suc k) t = + R.catchTC + (try1g t) + (R.catchTC (tryInv k t) + (R.catchTC (tryOp k t) (atom t))) + + + +compareTerms : R.Term → R.Term → R.TC Bool +compareTerms x x' = + ((R.noConstraints $ R.unify x x') >> (R.returnTC true)) <|> + (R.returnTC false) + + + +lookupOrAppend : List R.Term → R.Term → R.TC ((List R.Term) × ℕ) +lookupOrAppend [] t = R.returnTC ([ t ] , 0) +lookupOrAppend xs@(x ∷ xs') t = do + x≟t ← compareTerms t x + if x≟t + then (R.returnTC (xs , 0)) + else (do (xs'' , k) ← lookupOrAppend xs' t + R.returnTC (x ∷ xs'' , suc k)) + + +wildCatSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) +wildCatSolverTerm debugFlag t-g hole = do + + (t0 , t1) ← R.withNormalisation true $ + R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + (r0') ← tryGE.tryG t-g 100 t0 + (r1') ← tryGE.tryG t-g 100 t1 + (tmL , tmE0) ← ExprAccumM lookupOrAppend [] r0' + (tmL , tmE1) ← ExprAccumM lookupOrAppend tmL r1' + + let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) + pa = λ tmE → + let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) + rLpaTm = foldl + (λ x y → + (R∙ x ( R.def (quote WGTerm.ev[_]₂≡) + (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) + Rrefl rL + in ((R.def (quote _∙_) + (R.def (quote WGTerm.ev⤋²) + (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ []) + v∷ rLpaTm v∷ [] )) , + (Li.foldr + (λ x → Expr.showN ℕ true mkNiceVar x ∷nl_ ) [] $ rL)) + + let final = (R.def (quote _∙∙_∙∙_) + (fst (pa tmE0) + v∷ R.def (quote refl) [] + v∷ R.def (quote sym) (fst (pa tmE1) v∷ []) v∷ [] )) + + + let info = snd (pa tmE0) ++ snd (pa tmE1) ++ (Expr.showN ℕ true mkNiceVar tmE0 + ∷nl Expr.showN ℕ true mkNiceVar tmE1 + ∷nl niceAtomList tmL) + + R.returnTC (final , info) + + +-- wildCatSolverMain : Bool → R.Term → R.Term → R.TC Unit +-- wildCatSolverMain debugFlag t-g hole = do +-- ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type +-- hole' ← R.withNormalisation true $ R.checkType hole ty +-- (solution , msg) ← groupoidSolverTerm debugFlag t-g hole' +-- R.catchTC +-- (R.unify hole solution) +-- (R.typeError msg) + + +groupoidSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag t-g hole = do + + (t0 , t1) ← R.withNormalisation true $ + R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + (r0') ← tryGE.tryG t-g 100 t0 + (r1') ← tryGE.tryG t-g 100 t1 + (tmL , tmE0) ← ExprAccumM lookupOrAppend [] r0' + (tmL , tmE1) ← ExprAccumM lookupOrAppend tmL r1' + + let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) + pa = λ tmE → + let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) + rLpaTm = foldl + (λ x y → + (R∙ x ( R.def (quote WGTerm.ev[_]₂≡) + (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) + Rrefl rL + in ((R.def (quote _∙_) + (R.def (quote WGTerm.ev⤋²) + (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ []) + v∷ rLpaTm v∷ [] )) , + (Li.foldr + (λ x → Expr.showN ℕ true mkNiceVar x ∷nl_ ) [] $ rL)) + + let final = (R.def (quote _∙∙_∙∙_) + (fst (pa tmE0) + v∷ R.def (quote refl) [] + v∷ R.def (quote sym) (fst (pa tmE1) v∷ []) v∷ [] )) + + + let info = snd (pa tmE0) ++ snd (pa tmE1) ++ (Expr.showN ℕ true mkNiceVar tmE0 + ∷nl Expr.showN ℕ true mkNiceVar tmE1 + ∷nl niceAtomList tmL) + + R.returnTC (final , info) + + +groupoidSolverMain : Bool → R.Term → R.Term → R.TC Unit +groupoidSolverMain debugFlag t-g hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag t-g hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + + + + +macro + solveGroupoid : R.Term → R.Term → R.TC Unit + solveGroupoid = groupoidSolverMain true + + + + diff --git a/Cubical/WildCat/Base.agda b/Cubical/WildCat/Base.agda index fca6167fe7..66a4f3e7b4 100644 --- a/Cubical/WildCat/Base.agda +++ b/Cubical/WildCat/Base.agda @@ -8,6 +8,7 @@ module Cubical.WildCat.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function renaming (_∘_ to _∘f_) open import Cubical.Data.Sigma renaming (_×_ to _×'_) @@ -86,3 +87,55 @@ _^op : WildCat ℓ ℓ' → WildCat ℓ ℓ' (C ^op) .⋆IdL = C .⋆IdR (C ^op) .⋆IdR = C .⋆IdL (C ^op) .⋆Assoc f g h = sym (C .⋆Assoc _ _ _) + +IsWildGroupoid : (wc : WildCat ℓ ℓ') → Type (ℓ-max ℓ ℓ') +IsWildGroupoid wc = ∀ {x y} f → wildIsIso {C = wc} {x} {y} f + + +record WildGroupoid ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + no-eta-equality + + field + wildCat : WildCat ℓ ℓ' + isWildGroupoid : IsWildGroupoid wildCat + + _⋆'_ = _⋆_ wildCat + + inv : ∀ {x y} → wildCat [ x , y ] → wildCat [ y , x ] + inv f = wildIsIso.inv' (isWildGroupoid f) + + ⋆InvR : ∀ {x y} → (f : wildCat [ x , y ]) → f ⋆' inv f ≡ id wildCat + ⋆InvR f = wildIsIso.retr (isWildGroupoid f) + + ⋆InvL : ∀ {x y} → (f : wildCat [ x , y ]) → inv f ⋆' f ≡ id wildCat + ⋆InvL f = wildIsIso.sect (isWildGroupoid f) + + + invol-inv : ∀ {x y} → (f : wildCat [ x , y ]) → inv (inv f) ≡ f + invol-inv f = + (sym (⋆IdL wildCat (inv (inv f))) ∙ cong (_⋆' (inv (inv f))) (sym (⋆InvR _)) ) + ∙ ⋆Assoc wildCat _ _ _ ∙ cong (f ⋆'_) (⋆InvR (inv f)) + ∙ ⋆IdR wildCat f + + swapInv : ∀ {x y} → (f : wildCat [ x , y ]) → ∀ g → (inv f) ≡ g → f ≡ inv g + swapInv f g p = + sym (invol-inv f) ∙ cong inv p + + invUniqueR : ∀ {x y} {g : wildCat [ x , y ]} {h} → g ⋆' h ≡ id wildCat → h ≡ inv g + invUniqueR {g = g} {h} p = + (sym (⋆IdL wildCat h) ∙∙ cong (_⋆' h) (sym (⋆InvL g)) + ∙∙ ⋆Assoc wildCat (inv g) g h) ∙∙ cong ((inv g) ⋆'_) p ∙∙ ⋆IdR wildCat (inv g) + + + distInv : ∀ {x y z} (f : wildCat [ x , y ]) (g : wildCat [ y , z ]) → + inv (f ⋆' g) ≡ inv g ⋆' inv f + distInv f g = sym $ invUniqueR $ + (sym (⋆Assoc wildCat _ _ _) ∙∙ cong (_⋆' inv f) (⋆Assoc wildCat _ _ _) + ∙∙ (λ i → ⋆Assoc wildCat f ((⋆InvR g) i) (inv f) i)) + ∙∙ cong (f ⋆'_) (⋆IdL wildCat _) ∙∙ ⋆InvR f + + id≡inv-id : ∀ {x} → inv (id wildCat) ≡ id wildCat {x = x} + id≡inv-id = sym (⋆IdL wildCat (inv (id wildCat))) ∙ ⋆InvR (id wildCat) + + + open WildCat wildCat public diff --git a/Cubical/WildCat/Functor.agda b/Cubical/WildCat/Functor.agda index 352a3a0fff..05b363b02c 100644 --- a/Cubical/WildCat/Functor.agda +++ b/Cubical/WildCat/Functor.agda @@ -2,6 +2,7 @@ module Cubical.WildCat.Functor where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Data.Sigma using (ΣPathP) @@ -162,3 +163,13 @@ WildFunctor.F-ob commFunctor (x , y) = y , x WildFunctor.F-hom commFunctor (f , g) = g , f WildFunctor.F-id commFunctor = refl WildFunctor.F-seq commFunctor _ _ = refl + +module _ + {C : WildGroupoid ℓC ℓC'} {D : WildGroupoid ℓD ℓD'} + (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where + + module gC = WildGroupoid C + module gD = WildGroupoid D + + F-inv : ∀ {x y} f → F-hom F (gC.inv {x} {y} f) ≡ gD.inv (F-hom F f) + F-inv f = gD.invUniqueR $ sym (F-seq F _ _) ∙∙ congS (F-hom F) (gC.⋆InvR f) ∙∙ F-id F From 37aef30663cd4cf93b144dd3d2a7247c4c93139e Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Wed, 13 Mar 2024 03:40:01 +0100 Subject: [PATCH 02/20] wild cat solver --- Cubical/Tactics/GroupoidSolver/Example.agda | 41 +- Cubical/Tactics/GroupoidSolver/Solver.agda | 515 +++++++++----------- Cubical/WildCat/Base.agda | 2 +- 3 files changed, 266 insertions(+), 292 deletions(-) diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index d0a19509cd..5f32b7f1d9 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -6,15 +6,48 @@ open import Cubical.Foundations.Prelude open import Cubical.WildCat.Base open import Cubical.Tactics.GroupoidSolver.Solver +open import Cubical.Data.Maybe private variable ℓ ℓ' : Level + +module exampleC (WC : WildCat ℓ ℓ') where + -- open WildGroupoid WG + open WildCat WC + + WC' : WildStr ℓ ℓ' + WC' = WC , nothing + + module _ (x y z w v : ob) (p : Hom[ x , y ]) (q : Hom[ y , z ]) + (r : Hom[ z , w ]) (s : Hom[ w , v ]) where + + + pA pB pC : Hom[ x , v ] + pA = (p ⋆ (id ⋆ q)) ⋆ (r ⋆ s) + pB = (p ⋆ (q ⋆ ((r ⋆ id) ⋆ s))) ⋆ id + pC = ((p ⋆ (id ⋆ q)) ⋆ r) ⋆ s + + + pA=pB : pA ≡ pB + pA=pB = solveWildCat WC' + + -- pA=pC : pA ≡ pC + -- pA=pC = solveGroupoid WG' + + -- pB=pC : pB ≡ pC + -- pB=pC = solveGroupoid WG' + + module example (WG : WildGroupoid ℓ ℓ') where open WildGroupoid WG - + open WildCat wildCat + + WG' : WildStr ℓ ℓ' + WG' = wildCat , just isWildGroupoid + module _ (x y z w v : ob) (p p' : Hom[ x , y ]) (q q' : Hom[ y , z ]) (r : Hom[ w , z ]) (s : Hom[ w , v ]) where @@ -26,10 +59,10 @@ module example (WG : WildGroupoid ℓ ℓ') where pA=pB : pA ≡ pB - pA=pB = solveGroupoid WG + pA=pB = solveGroupoid WG' pA=pC : pA ≡ pC - pA=pC = solveGroupoid WG + pA=pC = solveGroupoid WG' pB=pC : pB ≡ pC - pB=pC = solveGroupoid WG + pB=pC = solveGroupoid WG' diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index 2c16c3ef51..50cd30d1c9 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -198,244 +198,28 @@ module Nodes (ob : Type ℓ) (Hom[_,_] : ob → ob → Type ℓ') (hasInvs : Boo no (red[ suc (suc k) , x ] ∷N x₁) red[ suc (suc k) , no (involN x x₁) ] = red[ suc (suc k) , x ] - - - module Ev (id : isRefl') - (_⋆_ : isTrans') - (inv : {_ : Bool→Type hasInvs} → isSym') where - - - eva[_] : ∀ {k} → ∀ {x y} → Atom k x y → Hom[ x , y ] - eva[ a⟦ x ⟧ ] = x - eva[ a⟦ x ⟧⁻ {invGuard} ] = inv {invGuard} x - - ev[_] : ∀ {k} → ∀ {x y} → Node k x y → Hom[ x , y ] - ev[_]b : ∀ {k} → ∀ {x y} → Node k x y → Bool → Hom[ x , y ] - - ev[ f ] = ev[ f ]b false - - - ev[ idN ]b _ = id - ev[ atomN x ]b _ = eva[ x ] - ev[ no (xs ∷N x) ]b b = ev[ xs ]b b ⋆ eva[ x ] - ev[ no (x ⋆N x₁) ]b b = ev[ x ]b b ⋆ ev[ x₁ ]b b - ev[ no (invN x {invGuard}) ]b b = inv {invGuard} (ev[ x ]b b) - ev[ no (involN x x₁ {invGuard}) ]b false = (ev[ x ] ⋆ eva[ x₁ ]) ⋆ - eva[ invAtom _ {1} invGuard x₁ ] - ev[ no (involN x x₁) ]b true = ev[ x ]b true - - module Ev' (⋆Assoc : ∀ {u v w x} - (f : Hom[ u , v ]) - (g : Hom[ v , w ]) - (h : Hom[ w , x ]) - → (f ⋆ g) ⋆ h ≡ f ⋆ (g ⋆ h)) - (⋆IdR : ∀ {x y} (f : Hom[ x , y ]) → f ⋆ id ≡ f) - (⋆IdL : ∀ {x y} (f : Hom[ x , y ]) → id ⋆ f ≡ f) - (⋆InvR : ∀ {x y hi} → (f : Hom[ x , y ]) → f ⋆ inv {hi} f ≡ id) - (⋆InvL : ∀ {x y hi} → (f : Hom[ x , y ]) → inv {hi} f ⋆ f ≡ id) - where - - evInv : ∀ {k} → ∀ {x y hi} → (f : Atom k x y) → - id ≡ (eva[ f ] ⋆ eva[ invAtom k {1} hi f ]) - evInv a⟦ x ⟧ = sym (⋆InvR x) - evInv a⟦ x ⟧⁻ = sym (⋆InvL x) - - ev[_]≡ : ∀ {k} → ∀ {x y} → (f : Node k x y) → ev[ f ]b false ≡ ev[ f ]b true - ev[ idN ]≡ = refl - ev[ atomN x ]≡ = refl - ev[ no (x ∷N x₁) ]≡ = cong (_⋆ eva[ x₁ ]) ev[ x ]≡ - ev[ no (x ⋆N x₁) ]≡ = cong₂ _⋆_ ev[ x ]≡ ev[ x₁ ]≡ - ev[ no (invN x) ]≡ = cong inv ev[ x ]≡ - ev[ no (involN x x₁) ]≡ = - ⋆Assoc _ _ _ - ∙∙ cong₂ _⋆_ ev[ x ]≡ (sym (evInv x₁)) - ∙∙ ⋆IdR _ - - ev[_]₂≡ = ev[_]≡ {k = 2} - - - evN2++b : ∀ {k} {a₀ a₁ a₂ : ob} b → (x : Node (suc (suc k)) a₀ a₁) - → (x₁ : Node (suc (suc k)) a₁ a₂) - → ev[ x ]b b ⋆ ev[ x₁ ]b b ≡ ev[ x N2++ x₁ ]b b - evN2++b b x idN = ⋆IdR _ - evN2++b b x (no (x₁ ∷N x₂)) = - sym (⋆Assoc _ _ _) - ∙ cong (_⋆ eva[ x₂ ]) (evN2++b b x x₁) - evN2++b false x (no (involN x₁ x₂)) = - sym (⋆Assoc _ _ _) - ∙∙ congS (_⋆ eva[ invAtom _ _ x₂ ]) (sym (⋆Assoc _ _ _)) - ∙∙ congS (λ y → ((y ⋆ eva[ x₂ ]) ⋆ eva[ invAtom _ _ x₂ ])) (evN2++b false x x₁) - - evN2++b true x (no (involN x₁ x₂)) = evN2++b true x x₁ - - evN2++ : ∀ {k} {a₀ a₁ a₂ : ob} → (x : Node (suc (suc k)) a₀ a₁) - → (x₁ : Node (suc (suc k)) a₁ a₂) - → ev[ x ] ⋆ ev[ x₁ ] ≡ ev[ x N2++ x₁ ] - evN2++ = evN2++b false - - module Ev'' (id≡inv-id : ∀ {x hi} → id {x} ≡ inv {hi} id) - (involInv : ∀ {x y hi hi'} → (f : Hom[ x , y ]) → - inv {hi} (inv {hi'} f) ≡ f) - (distInv : ∀ {x y z hi} - (f : Hom[ x , y ]) - (g : Hom[ y , z ]) - → inv {hi} (f ⋆ g) ≡ inv {hi} g ⋆ inv {hi} f) - where - - inv-eva : ∀ {x y} {k} {k'} a {hi} → - inv {hi} {x} {y} (eva[_] {k} a) ≡ eva[ invAtom (k) {k'} hi a ] - inv-eva a⟦ x ⟧ = refl - inv-eva a⟦ x ⟧⁻ {hi} = involInv x - - invAtomInvol : ∀ {x y} k {hi hi'} → (x : Atom (suc (suc k)) x y) → - (eva[ invAtom 2 {suc k} (hi) (invAtom (suc (suc k)) {1} hi' x) ]) ≡ (eva[ x ]) - invAtomInvol k a⟦ x ⟧ = refl - invAtomInvol k {hi} (a⟦ x ⟧⁻ {hi'}) i = inv {isPropBool→Type hi hi' i} x - - - ev[_]≡inv : ∀ {k} → ∀ {x y} {hi} b → (f : Node k x y) → - inv {hi} (ev[ f ]b b) ≡ ev[ invNode* {hi = hi} f ]b b - ev[_]≡inv {zero} b f = refl - ev[_]≡inv {suc k} {hi = hi} b idN = sym (id≡inv-id {hi = hi}) - ev[_]≡inv {suc zero} {hi = hi} b (atomN x) = inv-eva x {hi} - ev[_]≡inv {suc zero} b (no (x ⋆N x₁)) = - distInv (ev[ x ]b b) (ev[ x₁ ]b b) ∙ - cong₂ _⋆_ (ev[_]≡inv b x₁) (ev[_]≡inv b x) - ev[_]≡inv {suc (suc k)} {hi = hi} b (no (x ∷N x₁)) = - distInv (ev[ x ]b b) (eva[ x₁ ]) - ∙∙ cong₂ _⋆_ (λ i → ⋆IdL (inv-eva {k' = suc k} x₁ {hi} i) (~ i)) - (ev[_]≡inv b x) - ∙∙ evN2++b b (no (idN ∷N invAtom (suc (suc k)) hi x₁)) (invNode x) - - ev[_]≡inv {suc (suc k)} {hi = hi} false (no (involN x x₁ {hi'})) = - distInv (ev[ x ] ⋆ eva[ x₁ ]) (eva[ invAtom (suc (suc k)) hi' x₁ ]) - ∙∙ - cong₂ _⋆_ - (inv-eva {k' = suc k} (invAtom (suc (suc k)) hi' x₁)) - (distInv ev[ x ] eva[ x₁ ]) ∙ - (λ i → ⋆Assoc (⋆IdL (invAtomInvol k {hi} {hi'} x₁ i) (~ i)) - (inv-eva {k' = 1} x₁ {isPropBool→Type hi hi' i} i) - (ev[_]≡inv {suc (suc k)} {hi = hi} false x i) (~ i) ) - ∙∙ evN2++b false (no (involN idN x₁)) (invNode x) - - ev[_]≡inv {suc (suc k)} true (no (involN x x₁)) = - ev[_]≡inv true x - ∙∙ sym (⋆IdL _) - ∙∙ evN2++b true (no (involN idN x₁)) (invNode x) - - - - - data NodeCase : {a₀ a₁ : ob} → Hom[ a₀ , a₁ ] → Type (ℓ-max ℓ ℓ') where - idCase : ∀ {x} → NodeCase (id {a = x}) - opCase : ∀ {x y z : _} → (p : Hom[ x , y ]) (q : Hom[ y , z ]) → NodeCase (p ⋆ q) - invCase : ∀ {x y : _} → {hi : Bool→Type hasInvs} (p : Hom[ x , y ]) → NodeCase (inv {hi} p) - - -module _ (WC : WildCat ℓ ℓ') where - - open WildCat WC - - module WGi hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) {hi : Bool→Type hasInvs} where - open WildGroupoid (record { wildCat = WC ; isWildGroupoid = iwg hi }) public - using (inv ; ⋆InvR; ⋆InvL; id≡inv-id; distInv; invol-inv) - - invol-inv' : ∀ hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) - {hi hi' : Bool→Type hasInvs} {x y : WC .WildCat.ob} - (f : Hom[ x , y ]) → - WildGroupoid.inv - (record { wildCat = WC ; isWildGroupoid = iwg hi }) - (WildGroupoid.inv - (record { wildCat = WC ; isWildGroupoid = iwg hi' }) f) - ≡ f - invol-inv' true iwg {hi} {hi'} = WGi.invol-inv true iwg - - module WCTerm hasInvs (iwg : Bool→Type hasInvs → IsWildGroupoid WC) {hi : Bool→Type hasInvs} where - - open WGi hasInvs iwg - - open Nodes ob Hom[_,_] hasInvs public - open Ev id _⋆_ (λ {ig} → wildIsIso.inv' Fu.∘ iwg ig) public - - - module WEv' = Ev' ⋆Assoc ⋆IdR ⋆IdL ⋆InvR ⋆InvL - - - module WEv'' = WEv'.Ev'' (sym (id≡inv-id)) (invol-inv' hasInvs iwg) distInv - - - eva⤋ : ∀ k {a₀ a₁ : ob} → ∀ (h : Atom k a₀ a₁) → eva[ h ] ≡ eva[ k a⤋ h ] - eva⤋ k a⟦ x ⟧ = refl - eva⤋ k a⟦ x ⟧⁻ = refl - - - invAtom⤋ : ∀ k k' {hi} {a₀ a₁ : ob} → ∀ (h : Atom (suc (suc k)) a₀ a₁) → - eva[ invAtom (suc (suc k)) {k'} hi h ] - ≡ eva[ invAtom (suc (suc (suc k))) {k'} hi (suc (suc k) a⤋ h) ] - invAtom⤋ k k' Nodes.a⟦ x ⟧ = refl - invAtom⤋ k k' Nodes.a⟦ x ⟧⁻ = refl - - - ev⤋' : ∀ k {a₀ a₁ : ob} → ∀ (f' : Node' k a₀ a₁) → ev[ no f' ] ≡ ev[ k ⤋' f' ] - ev⤋ : ∀ k {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ k ⤋ f ] - - - - ev⤋ k idN = refl - ev⤋ zero (Nodes.atomN Nodes.a⟦ x ⟧) = refl - ev⤋ (suc zero) (Nodes.atomN Nodes.a⟦ x ⟧) = sym (⋆IdL x) - ev⤋ (suc zero) (Nodes.atomN (Nodes.a⟦ x ⟧⁻)) = sym ((⋆IdL _)) - ev⤋ k (no x) = ev⤋' k x - - ev⤋' (suc (suc k)) (x ∷N x₁) = - cong₂ _⋆_ (ev⤋ (suc (suc k)) x) (eva⤋ (suc (suc k)) x₁ ) - ev⤋' zero (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ zero x) (ev⤋ zero x₁) - ev⤋' (suc zero) (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ 1 x) (ev⤋ 1 x₁) ∙ - WEv'.evN2++ (1 ⤋ x) (1 ⤋ x₁) - ev⤋' zero (invN x {hi}) = - cong inv (ev⤋ zero x) ∙ WEv''.ev[_]≡inv false (0 ⤋ x) -- enInv1Node x hi - ev⤋' (suc (suc k)) (Nodes.involN x x₁ {hi}) = - cong₂ _⋆_ (λ i → (ev⤋ (suc (suc k)) x i ⋆ (eva⤋ (suc (suc k)) x₁) i)) - (invAtom⤋ k 1 {hi} x₁) - - ev⤋⁺ : ∀ {k} m {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ m ⤋⁺ f ] - ev⤋⁺ zero f = refl - ev⤋⁺ (suc m) f = ev⤋⁺ m f ∙ ev⤋ (m + _) (m ⤋⁺ f) - - ev⤋² = ev⤋⁺ {0} 2 - - - -module WGTerm (WG : WildGroupoid ℓ ℓ') where - open WCTerm (WildGroupoid.wildCat WG) true (λ _ → WildGroupoid.isWildGroupoid WG) public - - open WGi (WildGroupoid.wildCat WG) true (λ _ → WildGroupoid.isWildGroupoid WG) - - open WildGroupoid WG hiding (⋆InvR ; ⋆InvL) - open Ev' ⋆Assoc ⋆IdR ⋆IdL ⋆InvR ⋆InvL public - module _ (A : Type ℓ) where - module Expr = Nodes Unit (λ _ _ → A) + module Expr = Nodes Unit (λ _ _ → A) - module DecNodes (_≟A_ : Discrete A) where + module DecNodes (_≟A_ : Discrete A) where - AtomTo𝟚×A : Expr.Atom true 2 _ _ → (Bool × A) - AtomTo𝟚×A Nodes.a⟦ x ⟧ = true , x - AtomTo𝟚×A Nodes.a⟦ x ⟧⁻ = false , x + AtomTo𝟚×A : Expr.Atom true 2 _ _ → (Bool × A) + AtomTo𝟚×A Nodes.a⟦ x ⟧ = true , x + AtomTo𝟚×A Nodes.a⟦ x ⟧⁻ = false , x - mbRed : Expr.Node true 2 _ _ → Maybe (Expr.Node true 2 _ _) - mbRed Nodes.idN = nothing - mbRed (Nodes.no (Nodes.idN Nodes.∷N x₁)) = nothing - mbRed (Nodes.no (x'@(Nodes.no (x Nodes.∷N x₂)) Nodes.∷N x₁)) = - decRec (λ _ → just $ Nodes.no (Nodes.involN (Mb.rec x (idfun _) (mbRed x)) x₂) ) - (λ _ → map-Maybe (Nodes.no ∘ Nodes._∷N x₁) (mbRed x')) - (discreteΣ 𝟚._≟_ (λ _ → _≟A_) (AtomTo𝟚×A x₂) (AtomTo𝟚×A (Expr.invAtom true 2 _ x₁))) - mbRed (Nodes.no (Nodes.no (Nodes.involN x x₂) Nodes.∷N x₁)) = nothing - mbRed (Nodes.no (Nodes.involN x x₁)) = nothing + mbRed : Expr.Node true 2 _ _ → Maybe (Expr.Node true 2 _ _) + mbRed Nodes.idN = nothing + mbRed (Nodes.no (Nodes.idN Nodes.∷N x₁)) = nothing + mbRed (Nodes.no (x'@(Nodes.no (x Nodes.∷N x₂)) Nodes.∷N x₁)) = + decRec (λ _ → just $ Nodes.no (Nodes.involN (Mb.rec x (idfun _) (mbRed x)) x₂) ) + (λ _ → map-Maybe (Nodes.no ∘ Nodes._∷N x₁) (mbRed x')) + (discreteΣ 𝟚._≟_ (λ _ → _≟A_) (AtomTo𝟚×A x₂) (AtomTo𝟚×A (Expr.invAtom true 2 _ x₁))) + mbRed (Nodes.no (Nodes.no (Nodes.involN x x₂) Nodes.∷N x₁)) = nothing + mbRed (Nodes.no (Nodes.involN x x₁)) = nothing - redList : Expr.Node true 2 _ _ → List (Expr.Node true 2 _ _) - redList x = unfoldMaybe (Expr.len true x) (mbRed ∘ Expr.red[_,_] true 2) x + redList : Expr.Node true 2 _ _ → List (Expr.Node true 2 _ _) + redList x = unfoldMaybe (Expr.len true x) (mbRed ∘ Expr.red[_,_] true 2) x redListℕ = DecNodes.redList ℕ discreteℕ @@ -448,7 +232,7 @@ mapExpr {A = A} {A'} {b} {k} f = w wa : Expr.Atom A b k _ _ → Expr.Atom A' b k _ _ wa Nodes.a⟦ x ⟧ = Nodes.a⟦ f x ⟧ wa (Nodes.a⟦_⟧⁻ {invG = ig} x {g}) = Nodes.a⟦_⟧⁻ {invG = ig} (f x) {g} - + w : Expr.Node A b k _ _ → Expr.Node A' b k _ _ w Nodes.idN = Nodes.idN w (Nodes.atomN {aGuard = ag} x) = Nodes.atomN {aGuard = ag} (wa x) @@ -484,7 +268,7 @@ mapExprQ {A = A} {b} {k} f = w w Nodes.idN = R.con (quote Nodes.idN) [] w (Nodes.atomN x) = R.con (quote Nodes.atomN) (wa x v∷ []) w (Nodes.no x) = R.con (quote Nodes.no) (w' x v∷ []) - + ExprAccumM : ∀ {A : Type ℓ} {A' : Type ℓ'} {ℓs} {S : Type ℓs} {b k} → (S → A → R.TC (S × A')) → S @@ -497,7 +281,7 @@ ExprAccumM {A = A} {A'} {S = S} {b} {k} f = w wa : S → Expr.Atom A b k _ _ → R.TC (S × Expr.Atom A' b k _ _) wa s a⟦ x ⟧ = (λ (s' , x') → s' , a⟦ x' ⟧) <$> f s x wa s (a⟦_⟧⁻ {invG = g'} x {g}) = (λ (s' , x') → s' , (a⟦_⟧⁻ {invG = g'} x' {g})) <$> f s x - + w : S → Expr.Node A b k _ _ → R.TC (S × Expr.Node A' b k _ _) w' : S → Expr.Node' A b k _ _ → R.TC (S × Expr.Node A' b k _ _) w s idN = R.returnTC (s , idN) @@ -520,57 +304,196 @@ ExprAccumM {A = A} {A'} {S = S} {b} {k} f = w w' s (involN x x₁) = w s x -opCase' : ∀ (WG : WildGroupoid ℓ ℓ') {x y z} f g → - WGTerm.NodeCase WG {a₀ = x} {z} _ -opCase' WG {x} {y} {z} f g = WGTerm.opCase {WG = WG} {y = y} f g - -invCase' : ∀ (WG : WildGroupoid ℓ ℓ') {x y} f → - WGTerm.NodeCase WG {a₀ = y} {x} _ -invCase' WG {x} {y} f = WGTerm.invCase {WG = WG} {x = x} {y} f +WildStr : ∀ ℓ ℓ' → Type (ℓ-suc (ℓ-max ℓ ℓ')) +WildStr ℓ ℓ' = Σ (WildCat ℓ ℓ') (Maybe ∘ IsWildGroupoid) -id' : (WG : WildGroupoid ℓ ℓ') → ∀ {x} → WildGroupoid.Hom[_,_] WG x x -id' WG = WildGroupoid.id WG +hasInvs? : WildStr ℓ ℓ' → Bool +hasInvs? = caseMaybe false true ∘ snd + +WildStr→WG : (ws : WildStr ℓ ℓ') → Bool→Type (hasInvs? ws) → WildGroupoid ℓ ℓ' +WildGroupoid.wildCat (WildStr→WG (wc , _) _) = wc +WildGroupoid.isWildGroupoid (WildStr→WG (_ , just x) _) = x + + +module WSExpr (WS : WildStr ℓ ℓ') where + open WildCat (fst WS) public + open Nodes ob Hom[_,_] (hasInvs? WS) public + + InvGuard = Bool→Type (hasInvs? WS) + + module _ {ig : InvGuard} where + open WildGroupoid (WildStr→WG WS ig) public + + eva[_] : ∀ {k} → ∀ {x y} → Atom k x y → Hom[ x , y ] + eva[ a⟦ x ⟧ ] = x + eva[ a⟦ x ⟧⁻ {ig} ] = inv {ig} x + + ev[_] : ∀ {k} → ∀ {x y} → Node k x y → Hom[ x , y ] + ev[_]b : ∀ {k} → ∀ {x y} → Node k x y → Bool → Hom[ x , y ] + + ev[ f ] = ev[ f ]b false + + + ev[ idN ]b _ = id + ev[ atomN x ]b _ = eva[ x ] + ev[ no (xs ∷N x) ]b b = ev[ xs ]b b ⋆ eva[ x ] + ev[ no (x ⋆N x₁) ]b b = ev[ x ]b b ⋆ ev[ x₁ ]b b + ev[ no (invN x {invGuard}) ]b b = inv {invGuard} (ev[ x ]b b) + ev[ no (involN x x₁ {invGuard}) ]b false = (ev[ x ] ⋆ eva[ x₁ ]) ⋆ + eva[ invAtom _ {1} invGuard x₁ ] + ev[ no (involN x x₁) ]b true = ev[ x ]b true -inv' : (WG : WildGroupoid ℓ ℓ') → ∀ {x y} → WildGroupoid.Hom[_,_] WG x y → WildGroupoid.Hom[_,_] WG y x -inv' WG = WildGroupoid.inv WG + evInv : ∀ {k} → ∀ {x y hi} → (f : Atom k x y) → + id ≡ (eva[ f ] ⋆ eva[ invAtom k {1} hi f ]) + evInv a⟦ x ⟧ = sym (⋆InvR x) + evInv a⟦ x ⟧⁻ = sym (⋆InvL x) + ev[_]≡ : ∀ {k} → ∀ {x y} → (f : Node k x y) → ev[ f ]b false ≡ ev[ f ]b true + ev[ idN ]≡ = refl + ev[ atomN x ]≡ = refl + ev[ no (x ∷N x₁) ]≡ = cong (_⋆ eva[ x₁ ]) ev[ x ]≡ + ev[ no (x ⋆N x₁) ]≡ = cong₂ _⋆_ ev[ x ]≡ ev[ x₁ ]≡ + ev[ no (invN x) ]≡ = cong inv ev[ x ]≡ + ev[ no (involN x x₁) ]≡ = + ⋆Assoc _ _ _ + ∙∙ cong₂ _⋆_ ev[ x ]≡ (sym (evInv x₁)) + ∙∙ ⋆IdR _ -⋆' : (WG : WildGroupoid ℓ ℓ') → ∀ {x y z} → WildGroupoid.Hom[_,_] WG x y → WildGroupoid.Hom[_,_] WG y z → WildGroupoid.Hom[_,_] WG x z -⋆' WG = WildGroupoid._⋆_ WG + ev[_]₂≡ = ev[_]≡ {k = 2} + + + evN2++b : ∀ {k} {a₀ a₁ a₂ : ob} b → (x : Node (suc (suc k)) a₀ a₁) + → (x₁ : Node (suc (suc k)) a₁ a₂) + → ev[ x ]b b ⋆ ev[ x₁ ]b b ≡ ev[ x N2++ x₁ ]b b + evN2++b b x idN = ⋆IdR _ + evN2++b b x (no (x₁ ∷N x₂)) = + sym (⋆Assoc _ _ _) + ∙ cong (_⋆ eva[ x₂ ]) (evN2++b b x x₁) + evN2++b false x (no (involN x₁ x₂)) = + sym (⋆Assoc _ _ _) + ∙∙ congS (_⋆ eva[ invAtom _ _ x₂ ]) (sym (⋆Assoc _ _ _)) + ∙∙ congS (λ y → ((y ⋆ eva[ x₂ ]) ⋆ eva[ invAtom _ _ x₂ ])) (evN2++b false x x₁) + evN2++b true x (no (involN x₁ x₂)) = evN2++b true x x₁ -module ETerm = Expr R.Term + evN2++ : ∀ {k} {a₀ a₁ a₂ : ob} → (x : Node (suc (suc k)) a₀ a₁) + → (x₁ : Node (suc (suc k)) a₁ a₂) + → ev[ x ] ⋆ ev[ x₁ ] ≡ ev[ x N2++ x₁ ] + evN2++ = evN2++b false + + inv-eva : ∀ {x y} {k} {k'} a {hi} → + inv {hi} {x} {y} (eva[_] {k} a) ≡ eva[ invAtom (k) {k'} hi a ] + inv-eva a⟦ x ⟧ = refl + inv-eva (a⟦ x ⟧⁻ {hi}) {hi'} = + cong (inv {hi'}) ((λ i → inv {isPropBool→Type hi hi' i}) ≡$ x) ∙ invol-inv x + + invAtomInvol : ∀ {x y} k {hi hi'} → (x : Atom (suc (suc k)) x y) → + (eva[ invAtom 2 {suc k} (hi) (invAtom (suc (suc k)) {1} hi' x) ]) ≡ (eva[ x ]) + invAtomInvol k a⟦ x ⟧ = refl + invAtomInvol k {hi} (a⟦ x ⟧⁻ {hi'}) i = inv {isPropBool→Type hi hi' i} x -module _ (WGterm : R.Term) where - module EvTerm = ETerm.Ev true - (R.def (quote id') (WGterm v∷ [])) - (λ x y → (R.def (quote ⋆') (WGterm v∷ x v∷ y v∷ []))) - (λ x → (R.def (quote inv') (WGterm v∷ x v∷ []))) + + ev[_]≡inv : ∀ {k} → ∀ {x y} {hi} b → (f : Node k x y) → + inv {hi} (ev[ f ]b b) ≡ ev[ invNode* {hi = hi} f ]b b + ev[_]≡inv {zero} b f = refl + ev[_]≡inv {suc k} b idN = id≡inv-id + ev[_]≡inv {suc zero} {hi = hi} b (atomN x) = inv-eva x + ev[_]≡inv {suc zero} b (no (x ⋆N x₁)) = + distInv (ev[ x ]b b) (ev[ x₁ ]b b) ∙ + cong₂ _⋆_ (ev[_]≡inv b x₁) (ev[_]≡inv b x) + ev[_]≡inv {suc (suc k)} {hi = hi} b (no (x ∷N x₁)) = + distInv (ev[ x ]b b) (eva[ x₁ ]) + ∙∙ cong₂ _⋆_ (λ i → ⋆IdL (inv-eva {k' = suc k} x₁ {hi} i) (~ i)) + (ev[_]≡inv b x) + ∙∙ evN2++b b (no (idN ∷N invAtom (suc (suc k)) hi x₁)) (invNode x) + + ev[_]≡inv {suc (suc k)} {hi = hi} false (no (involN x x₁ {hi'})) = + distInv (ev[ x ] ⋆ eva[ x₁ ]) (eva[ invAtom (suc (suc k)) hi' x₁ ]) + ∙∙ + cong₂ _⋆_ + (inv-eva {k' = suc k} (invAtom (suc (suc k)) hi' x₁)) + (distInv ev[ x ] eva[ x₁ ]) ∙ + (λ i → ⋆Assoc (⋆IdL (invAtomInvol k {hi} {hi'} x₁ i) (~ i)) + (inv-eva {k' = 1} x₁ {isPropBool→Type hi hi' i} i) + (ev[_]≡inv {suc (suc k)} {hi = hi} false x i) (~ i) ) + ∙∙ evN2++b false (no (involN idN x₁)) (invNode x) + + ev[_]≡inv {suc (suc k)} true (no (involN x x₁)) = + ev[_]≡inv true x + ∙∙ sym (⋆IdL _) + ∙∙ evN2++b true (no (involN idN x₁)) (invNode x) + + + data NodeCase : {a₀ a₁ : ob} → Hom[ a₀ , a₁ ] → Type (ℓ-max ℓ ℓ') where + idCase : ∀ {x} → NodeCase (id {x = x}) + opCase : ∀ {x y z : _} → (p : Hom[ x , y ]) (q : Hom[ y , z ]) → NodeCase (p ⋆ q) + invCase : ∀ {x y : _} → {hi : InvGuard} (p : Hom[ x , y ]) → NodeCase (inv {hi} p) + + + eva⤋ : ∀ k {a₀ a₁ : ob} → ∀ (h : Atom k a₀ a₁) → eva[ h ] ≡ eva[ k a⤋ h ] + eva⤋ k a⟦ x ⟧ = refl + eva⤋ k a⟦ x ⟧⁻ = refl + + + invAtom⤋ : ∀ k k' {hi} {a₀ a₁ : ob} → ∀ (h : Atom (suc (suc k)) a₀ a₁) → + eva[ invAtom (suc (suc k)) {k'} hi h ] + ≡ eva[ invAtom (suc (suc (suc k))) {k'} hi (suc (suc k) a⤋ h) ] + invAtom⤋ k k' Nodes.a⟦ x ⟧ = refl + invAtom⤋ k k' Nodes.a⟦ x ⟧⁻ = refl + + + ev⤋' : ∀ k {a₀ a₁ : ob} → ∀ (f' : Node' k a₀ a₁) → ev[ no f' ] ≡ ev[ k ⤋' f' ] + ev⤋ : ∀ k {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ k ⤋ f ] + + + + ev⤋ k idN = refl + ev⤋ zero (Nodes.atomN Nodes.a⟦ x ⟧) = refl + ev⤋ (suc zero) (Nodes.atomN Nodes.a⟦ x ⟧) = sym (⋆IdL x) + ev⤋ (suc zero) (Nodes.atomN (Nodes.a⟦ x ⟧⁻)) = sym ((⋆IdL _)) + ev⤋ k (no x) = ev⤋' k x + + ev⤋' (suc (suc k)) (x ∷N x₁) = + cong₂ _⋆_ (ev⤋ (suc (suc k)) x) (eva⤋ (suc (suc k)) x₁ ) + ev⤋' zero (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ zero x) (ev⤋ zero x₁) + ev⤋' (suc zero) (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ 1 x) (ev⤋ 1 x₁) ∙ + evN2++ (1 ⤋ x) (1 ⤋ x₁) + ev⤋' zero (invN x {hi}) = + cong inv (ev⤋ zero x) ∙ ev[_]≡inv false (0 ⤋ x) -- enInv1Node x hi + ev⤋' (suc (suc k)) (Nodes.involN x x₁ {hi}) = + cong₂ _⋆_ (λ i → (ev⤋ (suc (suc k)) x i ⋆ (eva⤋ (suc (suc k)) x₁) i)) + (invAtom⤋ k 1 {hi} x₁) + + ev⤋⁺ : ∀ {k} m {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ m ⤋⁺ f ] + ev⤋⁺ zero f = refl + ev⤋⁺ (suc m) f = ev⤋⁺ m f ∙ ev⤋ (m + _) (m ⤋⁺ f) + + ev⤋² = ev⤋⁺ {0} 2 module Eℕ = Expr ℕ true +module ETerm = Expr R.Term NodeTerm : Bool → ℕ → Type ℓ-zero NodeTerm = λ b k → Expr.Node R.Term b k tt tt +module tryWCE (tG : R.Term) where -module tryGE (tG : R.Term) where - tryG : ℕ → R.Term → R.TC (NodeTerm true 0) try1g : R.Term → R.TC (NodeTerm true 0) try1g t = do - _ ← R.unify t (R.def (quote id') [ varg tG ]) + _ ← R.unify t (R.def (quote WSExpr.id) [ varg tG ]) R.returnTC (ETerm.idN) tryOp : ℕ → R.Term → R.TC (NodeTerm true 0) tryOp zero _ = R.typeError [] tryOp (suc k) t = do - tm ← R.withNormalisation true $ R.checkType (R.def (quote opCase') - (varg tG ∷ varg R.unknown ∷ [ varg R.unknown ])) - (R.def (quote WGTerm.NodeCase) ((varg tG) ∷ [ varg t ])) + tm ← R.withNormalisation true $ R.checkType (R.con (quote WSExpr.opCase) + (varg R.unknown ∷ [ varg R.unknown ])) + (R.def (quote WSExpr.NodeCase) ((varg tG) ∷ [ varg t ])) (t1 , t2) ← h tm t1' ← tryG k t1 t2' ← tryG k t2 @@ -586,8 +509,8 @@ module tryGE (tG : R.Term) where tryInv zero _ = R.typeError [] tryInv (suc k) t = do tm ← R.withNormalisation true $ - (R.checkType (R.def (quote invCase') - ((varg tG) ∷ [ varg R.unknown ])) (R.def (quote WGTerm.NodeCase) + (R.checkType (R.con (quote WSExpr.invCase) + ([ varg R.unknown ])) (R.def (quote WSExpr.NodeCase) ((varg tG) ∷ [ varg t ]))) R.debugPrint "tryInv" 30 ([ R.strErr "\n ---- \n" ]) R.debugPrint "tryInv" 30 ([ R.termErr t ]) @@ -620,8 +543,12 @@ module tryGE (tG : R.Term) where R.catchTC (try1g t) (R.catchTC (tryInv k t) - (R.catchTC (tryOp k t) (atom t))) - + -- (tryOp k t) + (R.catchTC (tryOp k t) (atom t)) + ) + + + compareTerms : R.Term → R.Term → R.TC Bool @@ -641,6 +568,7 @@ lookupOrAppend xs@(x ∷ xs') t = do R.returnTC (x ∷ xs'' , suc k)) + wildCatSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) wildCatSolverTerm debugFlag t-g hole = do @@ -649,23 +577,19 @@ wildCatSolverTerm debugFlag t-g hole = do (R.typeError [ R.strErr "unable to get boundary" ]) (λ x → R.returnTC x) - (r0') ← tryGE.tryG t-g 100 t0 - (r1') ← tryGE.tryG t-g 100 t1 - (tmL , tmE0) ← ExprAccumM lookupOrAppend [] r0' - (tmL , tmE1) ← ExprAccumM lookupOrAppend tmL r1' + (tmL , tmE0) ← tryWCE.tryG t-g 100 t0 >>= ExprAccumM lookupOrAppend [] + (tmL , tmE1) ← tryWCE.tryG t-g 100 t1 >>= ExprAccumM lookupOrAppend tmL let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) pa = λ tmE → let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) rLpaTm = foldl (λ x y → - (R∙ x ( R.def (quote WGTerm.ev[_]₂≡) + (R∙ x ( R.def (quote WSExpr.ev[_]₂≡) (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) Rrefl rL - in ((R.def (quote _∙_) - (R.def (quote WGTerm.ev⤋²) - (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ []) - v∷ rLpaTm v∷ [] )) , + in ((R.def (quote WSExpr.ev⤋²) + (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ [])) , (Li.foldr (λ x → Expr.showN ℕ true mkNiceVar x ∷nl_ ) [] $ rL)) @@ -682,14 +606,22 @@ wildCatSolverTerm debugFlag t-g hole = do R.returnTC (final , info) --- wildCatSolverMain : Bool → R.Term → R.Term → R.TC Unit --- wildCatSolverMain debugFlag t-g hole = do --- ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type --- hole' ← R.withNormalisation true $ R.checkType hole ty --- (solution , msg) ← groupoidSolverTerm debugFlag t-g hole' --- R.catchTC --- (R.unify hole solution) --- (R.typeError msg) +wildCatSolverMain : Bool → R.Term → R.Term → R.TC Unit +wildCatSolverMain debugFlag t-g hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← wildCatSolverTerm debugFlag t-g hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + + + + +macro + solveWildCat : R.Term → R.Term → R.TC Unit + solveWildCat = wildCatSolverMain true + groupoidSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) @@ -700,21 +632,19 @@ groupoidSolverTerm debugFlag t-g hole = do (R.typeError [ R.strErr "unable to get boundary" ]) (λ x → R.returnTC x) - (r0') ← tryGE.tryG t-g 100 t0 - (r1') ← tryGE.tryG t-g 100 t1 - (tmL , tmE0) ← ExprAccumM lookupOrAppend [] r0' - (tmL , tmE1) ← ExprAccumM lookupOrAppend tmL r1' + (tmL , tmE0) ← tryWCE.tryG t-g 100 t0 >>= ExprAccumM lookupOrAppend [] + (tmL , tmE1) ← tryWCE.tryG t-g 100 t1 >>= ExprAccumM lookupOrAppend tmL let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) pa = λ tmE → let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) rLpaTm = foldl (λ x y → - (R∙ x ( R.def (quote WGTerm.ev[_]₂≡) + (R∙ x ( R.def (quote WSExpr.ev[_]₂≡) (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) Rrefl rL in ((R.def (quote _∙_) - (R.def (quote WGTerm.ev⤋²) + (R.def (quote WSExpr.ev⤋²) (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ []) v∷ rLpaTm v∷ [] )) , (Li.foldr @@ -752,3 +682,14 @@ macro + + + + + + + + + + + diff --git a/Cubical/WildCat/Base.agda b/Cubical/WildCat/Base.agda index 66a4f3e7b4..fbaa83fc3a 100644 --- a/Cubical/WildCat/Base.agda +++ b/Cubical/WildCat/Base.agda @@ -138,4 +138,4 @@ record WildGroupoid ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where id≡inv-id = sym (⋆IdL wildCat (inv (id wildCat))) ∙ ⋆InvR (id wildCat) - open WildCat wildCat public + -- open WildCat wildCat public From e4bbc4d89afbfd1558d24d66f27cea7d4ebcb590 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 15:15:18 +0100 Subject: [PATCH 03/20] wip --- Cubical/Data/Bool/Properties.agda | 4 + Cubical/Data/List/Properties.agda | 8 + Cubical/Data/Maybe/Base.agda | 3 + Cubical/Data/Maybe/Properties.agda | 10 + Cubical/Data/Sigma/Properties.agda | 3 + Cubical/Reflection/Base.agda | 2 + Cubical/Tactics/GroupSolver/Example.agda | 36 - Cubical/Tactics/GroupSolver/Solver.agda | 44 - Cubical/Tactics/GroupoidSolver/Example.agda | 174 +++- Cubical/Tactics/GroupoidSolver/Solver.agda | 670 +------------- Cubical/Tactics/WildCatSolver/Example.agda | 170 ++++ .../Reflection.agda | 53 +- Cubical/Tactics/WildCatSolver/Solver.agda | 57 ++ Cubical/Tactics/WildCatSolver/Solvers.agda | 833 ++++++++++++++++++ Cubical/WildCat/Functor.agda | 34 +- 15 files changed, 1325 insertions(+), 776 deletions(-) delete mode 100644 Cubical/Tactics/GroupSolver/Example.agda delete mode 100644 Cubical/Tactics/GroupSolver/Solver.agda create mode 100644 Cubical/Tactics/WildCatSolver/Example.agda rename Cubical/Tactics/{GroupoidSolver => WildCatSolver}/Reflection.agda (77%) create mode 100644 Cubical/Tactics/WildCatSolver/Solver.agda create mode 100644 Cubical/Tactics/WildCatSolver/Solvers.agda diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index 5570fabbd6..a1d68bbaf0 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -324,6 +324,10 @@ Dec→DecBool (no ¬p) q = Empty.rec (¬p q) DecBool→Dec : {P : Type ℓ} → (dec : Dec P) → Bool→Type (Dec→Bool dec) → P DecBool→Dec (yes p) _ = p +Bool→Type-not-⊕ : ∀ {x y} → Bool→Type (not (x ⊕ y)) → Bool→Type x → Bool→Type y +Bool→Type-not-⊕ {false} {false} _ () +Bool→Type-not-⊕ {true} {true} _ = _ + Dec≃DecBool : {P : Type ℓ} → (h : isProp P)(dec : Dec P) → P ≃ Bool→Type (Dec→Bool dec) Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Type (Dec→DecBool dec) (DecBool→Dec dec) diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index 7fe83bfd03..50063cd2ab 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -184,3 +184,11 @@ lookupWithDefault : A → List A → ℕ → A lookupWithDefault a [] _ = a lookupWithDefault _ (x ∷ _) zero = x lookupWithDefault a (x ∷ xs) (suc k) = lookupWithDefault a xs k + +intersperse : A → List A → List A +intersperse _ [] = [] +intersperse a (x ∷ xs) = x ∷ a ∷ intersperse a xs + +join : List (List A) → List A +join [] = [] +join (x ∷ xs) = x ++ join xs diff --git a/Cubical/Data/Maybe/Base.agda b/Cubical/Data/Maybe/Base.agda index 1712b822ca..73505f3dda 100644 --- a/Cubical/Data/Maybe/Base.agda +++ b/Cubical/Data/Maybe/Base.agda @@ -28,3 +28,6 @@ rec n j (just a) = j a elim : ∀ {A : Type ℓA} (B : Maybe A → Type ℓB) → B nothing → ((x : A) → B (just x)) → (mx : Maybe A) → B mx elim B n j nothing = n elim B n j (just a) = j a + +fromMaybe : A → Maybe A → A +fromMaybe x = rec x (λ x → x) diff --git a/Cubical/Data/Maybe/Properties.agda b/Cubical/Data/Maybe/Properties.agda index b277f81e63..77b436af9f 100644 --- a/Cubical/Data/Maybe/Properties.agda +++ b/Cubical/Data/Maybe/Properties.agda @@ -179,3 +179,13 @@ congMaybeEquiv e = isoToEquiv isom isom .rightInv (just b) = cong just (secEq e b) isom .leftInv nothing = refl isom .leftInv (just a) = cong just (retEq e a) + +bind : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Maybe A → (A → Maybe B) → Maybe B +bind nothing _ = nothing +bind (just x) x₁ = x₁ x + +map2-Maybe : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → (A → B → C) → Maybe A → Maybe B → Maybe C +map2-Maybe _ nothing _ = nothing +map2-Maybe f (just x) = map-Maybe (f x) diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 7a1a0f53cd..7d8cdb2003 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -52,6 +52,9 @@ map-snd f (a , b) = (a , f b) map-× : {B : Type ℓ} {B' : Type ℓ'} → (A → A') → (B → B') → A × B → A' × B' map-× f g (a , b) = (f a , g b) +map-both : (A → A') → A × A → A' × A' +map-both f = map-× f f + ≡-× : {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → fst x ≡ fst y → snd x ≡ snd y → x ≡ y ≡-× p q i = (p i) , (q i) diff --git a/Cubical/Reflection/Base.agda b/Cubical/Reflection/Base.agda index 8b18ceba4f..dd81312c08 100644 --- a/Cubical/Reflection/Base.agda +++ b/Cubical/Reflection/Base.agda @@ -32,6 +32,8 @@ pattern harg {q = q} t = R.arg (R.arg-info R.hidden (R.modality R.relevant q)) t pattern _v∷_ a l = varg a ∷ l pattern _h∷_ a l = harg a ∷ l +pattern v[_] a = varg a ∷ [] + infixr 5 _v∷_ _h∷_ vlam : String → R.Term → R.Term diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda deleted file mode 100644 index ac032f53fb..0000000000 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ /dev/null @@ -1,36 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.GroupSolver.Example where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure - -open import Cubical.Algebra.Group -open import Cubical.Tactics.GroupSolver.Solver - -private - variable - ℓ : Level - - -module example (G : Group ℓ) where - open GroupStr (snd G) - - module _ (p p' q q' r s : ⟨ G ⟩) where - - - pA pB pC : ⟨ G ⟩ - pA = (p · (1g · q)) · inv r - pB = (p · ((q · (inv ((q' · inv q') · - inv r · ((s · inv s) · r)) · inv q)) · (q · 1g))) · (inv r · 1g) - pC = (1g · p') · (((q · inv q) · (inv p' · p)) · (q · (inv q · (q · inv r)))) - - - pA=pB : pA ≡ pB - pA=pB = solveGroup G - - pA=pC : pA ≡ pC - pA=pC = solveGroup G - - pB=pC : pB ≡ pC - pB=pC = solveGroup G diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda deleted file mode 100644 index a2bdaaa024..0000000000 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ /dev/null @@ -1,44 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.GroupSolver.Solver where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Unit -open import Cubical.Data.List as Li - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R - -open import Cubical.Tactics.Reflection -open import Cubical.WildCat.Base -open import Cubical.Tactics.GroupoidSolver.Solver - -open import Cubical.Algebra.Group - -private - variable - ℓ : Level - - -module _ (G : Group ℓ) where - open GroupStr (snd G) - Group→WildGroupoid : WildGroupoid ℓ-zero ℓ - WildCat.ob (WildGroupoid.wildCat Group→WildGroupoid) = Unit - WildCat.Hom[_,_] (WildGroupoid.wildCat Group→WildGroupoid) _ _ = ⟨ G ⟩ - WildCat.id (WildGroupoid.wildCat Group→WildGroupoid) = 1g - WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ - WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL - WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR - WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) - wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f - wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f - wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f - - -macro - solveGroup : R.Term → R.Term → R.TC Unit - solveGroup g-t = groupoidSolverMain true (R.def (quote Group→WildGroupoid ) (g-t v∷ [])) diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 5f32b7f1d9..8000876cfa 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -1,68 +1,170 @@ {-# OPTIONS --safe #-} -module Cubical.Tactics.GroupoidSolver.Example where +module Cubical.Tactics.WildCatSolver.Example where open import Cubical.Foundations.Prelude open import Cubical.WildCat.Base -open import Cubical.Tactics.GroupoidSolver.Solver -open import Cubical.Data.Maybe +open import Cubical.Tactics.WildCatSolver.Solver +open import Cubical.Data.List + +open import Cubical.WildCat.Functor private variable ℓ ℓ' : Level +module exampleC (WC WC* : WildCat ℓ ℓ') + (F : WildFunctor WC* WC) where - -module exampleC (WC : WildCat ℓ ℓ') where - -- open WildGroupoid WG open WildCat WC + module * = WildCat WC* - WC' : WildStr ℓ ℓ' - WC' = WC , nothing - - module _ (x y z w v : ob) (p : Hom[ x , y ]) (q : Hom[ y , z ]) - (r : Hom[ z , w ]) (s : Hom[ w , v ]) where + module _ x y z w v + (p : Hom[ x , F ⟅ y ⟆ ]) + (q : *.Hom[ y , z ]) + (r : *.Hom[ z , w ]) + (s : Hom[ F ⟅ w ⟆ , v ]) where pA pB pC : Hom[ x , v ] - pA = (p ⋆ (id ⋆ q)) ⋆ (r ⋆ s) - pB = (p ⋆ (q ⋆ ((r ⋆ id) ⋆ s))) ⋆ id - pC = ((p ⋆ (id ⋆ q)) ⋆ r) ⋆ s + pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) + pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) + pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) + open WildCat-Solver ℓ ℓ' pA=pB : pA ≡ pB - pA=pB = solveWildCat WC' + pA=pB = solveWildCat (WC ∷ WC* ∷ []) - -- pA=pC : pA ≡ pC - -- pA=pC = solveGroupoid WG' + pB=pC : pB ≡ pC + pB=pC = solveWildCat (WC ∷ WC* ∷ []) - -- pB=pC : pB ≡ pC - -- pB=pC = solveGroupoid WG' -module example (WG : WildGroupoid ℓ ℓ') where - open WildGroupoid WG - open WildCat wildCat +-- -- module example (WG WG* : WildGroupoid ℓ ℓ') +-- -- (F : WildFunctor +-- -- (WildGroupoid.wildCat WG*) +-- -- (WildGroupoid.wildCat WG)) +-- -- where - WG' : WildStr ℓ ℓ' - WG' = wildCat , just isWildGroupoid +-- -- module WildStrGPD = WildStr (GroupoidWS {ℓ} {ℓ'}) - module _ (x y z w v : ob) (p p' : Hom[ x , y ]) (q q' : Hom[ y , z ]) - (r : Hom[ w , z ]) (s : Hom[ w , v ]) where +-- -- -- open WildStrGPD using +-- -- -- (testInferExpr; solveNoInvs; solveW) +-- -- open WildGroupoid WG +-- -- open WildCat wildCat - pA pB pC : Hom[ x , w ] - pA = (p ⋆ (id ⋆ q)) ⋆ inv r - pB = (p ⋆ ((q ⋆ (inv (inv r ⋆ ((s ⋆ inv s) ⋆ r)) ⋆ inv q)) ⋆ (q ⋆ id))) ⋆ (inv r ⋆ id) - pC = (id ⋆ p') ⋆ (((q ⋆ inv q) ⋆ (inv p' ⋆ p)) ⋆ (q ⋆ (inv q ⋆ (q ⋆ inv r)))) +-- -- open WildFunctor - pA=pB : pA ≡ pB - pA=pB = solveGroupoid WG' +-- -- module * where +-- -- open WildCat (WildGroupoid.wildCat WG*) public +-- -- open WildGroupoid WG* public + +-- -- -- module T1 x y z w +-- -- -- (p : Hom[ x , y ]) +-- -- -- (q : Hom[ y , z ]) +-- -- -- (r : Hom[ z , w ]) where + + +-- -- -- pA pB : Hom[ x , w ] +-- -- -- pA = p ⋆ (q ⋆ r) +-- -- -- pB = (p ⋆ q) ⋆ r + +-- -- -- -- pF : + +-- -- -- pXX : Hom[ x , x ] +-- -- -- pXX = id + + +-- -- -- tryTestOp : Unit +-- -- -- tryTestOp = testTryOp GroupoidWS WG pA + +-- -- -- tryTestId : Unit +-- -- -- tryTestId = testTryId GroupoidWS WG pXX + +-- -- -- pACase : FuCases WG pA +-- -- -- pACase = p FuCases.⋆FE (q ⋆ r) + + +-- -- module T1 x y z w +-- -- (p p' : Hom[ x , y ]) +-- -- (q : Hom[ y , z ]) +-- -- (r : Hom[ z , w ]) where + + +-- -- -- pA pB : Hom[ x , x ] +-- -- -- pA = (p ⋆ inv p) +-- -- -- pB = id + +-- -- -- pA≡pB : pA ≡ pB +-- -- -- pA≡pB = {!solveW GroupoidWS (WG ∷ WG* ∷ [])!} + + +-- -- pA pB : Hom[ x , w ] +-- -- pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r +-- -- -- pA = ((((((id ⋆ p))) ⋆ (inv p)) ⋆ p) ⋆ q) ⋆ r +-- -- pB = ((id ⋆ p) ⋆ q) ⋆ r + +-- -- pA≡pB : pA ≡ pB +-- -- pA≡pB = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) + + + +-- -- module T2 x y z w +-- -- (p : Hom[ x , F ⟅ y ⟆ ]) +-- -- (p' : Hom[ F ⟅ y ⟆ , x ]) +-- -- (q : *.Hom[ z , y ]) +-- -- (r : *.Hom[ z , w ]) where + + +-- -- pA : Hom[ F ⟅ y ⟆ , F-ob F w ] +-- -- pA = F-hom F ((*.inv q) *.⋆ r) +-- -- pB pB' pB'' : Hom[ x , F ⟅ w ⟆ ] +-- -- pB = (p ⋆ (F-hom F ((*.inv q) *.⋆ r))) +-- -- pB'' = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) +-- -- pB' = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) + + +-- -- pB''≡pB' : pB'' ≡ pB' +-- -- pB''≡pB' = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) + + + + + + +-- -- -- module example (WG WG* : WildGroupoid ℓ ℓ') +-- -- -- (F : WildFunctor (WildGroupoid.wildCat WG*) (WildGroupoid.wildCat WG)) where +-- -- -- open WildGroupoid WG +-- -- -- open WildCat wildCat + + + +-- -- -- module * where +-- -- -- open WildCat (WildGroupoid.wildCat WG*) public +-- -- -- open WildGroupoid WG* public + +-- -- -- module _ x v y z w +-- -- -- (p p' : Hom[ x , F ⟅ y ⟆ ]) +-- -- -- (q q' : *.Hom[ y , z ]) +-- -- -- (r : *.Hom[ w , z ]) +-- -- -- (s : Hom[ F ⟅ w ⟆ , v ]) where + + +-- -- -- pA pB pC : Hom[ x , v ] +-- -- -- pA = (p ⋆ (id ⋆ F ⟪ q *.⋆ *.inv q' ⟫)) ⋆ (F ⟪ q' *.⋆ (*.inv r) ⟫ ⋆ s) +-- -- -- pB = ((p ⋆ F ⟪ q *.⋆ (*.inv (*.inv r *.⋆ r) *.⋆ *.inv q)⟫) +-- -- -- ⋆ F ⟪ q *.⋆ *.id ⟫) ⋆ ((inv (F ⟪ r ⟫ ⋆ id)) ⋆ s) + +-- -- -- pC = p' ⋆ (inv p' ⋆ (p ⋆ (F ⟪ q *.⋆ (*.inv r) ⟫ ⋆ s))) + +-- -- -- pA=pB : pA ≡ pB +-- -- -- pA=pB = solveGroupoid WG WG* + +-- -- -- pA=pC : pA ≡ pC +-- -- -- pA=pC = solveGroupoid WG WG* - pA=pC : pA ≡ pC - pA=pC = solveGroupoid WG' - pB=pC : pB ≡ pC - pB=pC = solveGroupoid WG' diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index 50cd30d1c9..f1c8259ece 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Tactics.GroupoidSolver.Solver where +module Cubical.Tactics.Groupoid.Solver where open import Cubical.Foundations.Prelude @@ -17,13 +17,14 @@ open import Cubical.Data.Nat.Order.Recursive open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ open import Cubical.Data.List as Li open import Cubical.Data.Maybe as Mb open import Cubical.HITs.Interval -open import Cubical.Relation.Nullary +-- open import Cubical.Relation.Nullary open import Cubical.Relation.Binary open import Cubical.Reflection.Base @@ -34,662 +35,23 @@ open import Agda.Builtin.String -- open import Cubical.WildCat.WGE open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor +open import Cubical.Algebra.Group +open import Cubical.Tactics.WildCatSolver.Solvers -private - variable - ℓ ℓ' : Level -unfoldMaybe : ℕ → ∀ {ℓ} {A : Type ℓ} → (A → Maybe A) → A → List A -unfoldMaybe zero _ _ = [] -unfoldMaybe (suc x) f a = - Mb.rec [] (λ a' → a' ∷ unfoldMaybe x f a') (f a) +module WildCat-Solver ℓ ℓ' where -data 𝑵𝒐𝒅𝒆 : Type where - 𝒐𝒑 𝒂𝒕𝒐𝒎 𝒄𝒐𝒏𝒔 𝒊𝒏𝒗 𝒊𝒏𝒗𝒐𝒍 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 : 𝑵𝒐𝒅𝒆 + WildCatWS : WildStr ℓ ℓ' + WildStr.wildStr (WildCatWS) = WildCat ℓ ℓ' + WildStr.toWildCat WildCatWS = idfun _ + WildStr.mbIsWildGroupoid WildCatWS = nothing -module Nodes (ob : Type ℓ) (Hom[_,_] : ob → ob → Type ℓ') (hasInvs : Bool) where + private + module WC-WS = WildStr WildCatWS - 𝒏 : ℕ → 𝑵𝒐𝒅𝒆 → Bool + macro + solveWildCat : R.Term → R.Term → R.TC Unit + solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) - 𝒏 0 𝒐𝒑 = true - 𝒏 0 𝒊𝒏𝒗 = hasInvs - - 𝒏 0 𝒄𝒐𝒏𝒔 = false - 𝒏 0 𝒂𝒕𝒐𝒎 = true - 𝒏 0 𝒊𝒏𝒗𝒐𝒍 = false - 𝒏 0 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 = false - - 𝒏 (suc _) 𝒊𝒏𝒗 = false - 𝒏 1 𝒐𝒑 = true - 𝒏 1 𝒄𝒐𝒏𝒔 = false - 𝒏 1 𝒂𝒕𝒐𝒎 = true - 𝒏 1 𝒊𝒏𝒗𝒐𝒍 = false - 𝒏 (suc _) 𝒊𝒏𝒗𝑨𝒕𝒐𝒎 = hasInvs - - 𝒏 (suc (suc _)) 𝒐𝒑 = false - 𝒏 (suc (suc _)) 𝒄𝒐𝒏𝒔 = true - 𝒏 (suc (suc _)) 𝒊𝒏𝒗𝒐𝒍 = hasInvs - 𝒏 (suc (suc _)) 𝒂𝒕𝒐𝒎 = false - - - open BinaryRelation Hom[_,_] public - - - module _ (k : ℕ) where - N : 𝑵𝒐𝒅𝒆 → Type - N = Bool→Type Fu.∘ 𝒏 k - - data Atom : ob → ob → Type (ℓ-max ℓ ℓ') where - a⟦_⟧ : ∀ {x y} → Hom[ x , y ] → Atom x y - a⟦_⟧⁻ : ∀ {x y} → {invG : N 𝒊𝒏𝒗𝑨𝒕𝒐𝒎} → Hom[ y , x ] → {Bool→Type hasInvs} → Atom x y - - Atom→𝟚×H : ∀ {x y} → Atom x y → Σ Bool λ { true → Hom[ x , y ] ; false → Hom[ y , x ] } - Atom→𝟚×H a⟦ x ⟧ = true , x - Atom→𝟚×H a⟦ x ⟧⁻ = false , x - - data Node' : ob → ob → Type (ℓ-max ℓ ℓ') - data Node : ob → ob → Type (ℓ-max ℓ ℓ') where - idN : ∀ {x} → Node x x - atomN : ∀ {x y} → {aGuard : N 𝒂𝒕𝒐𝒎} → Atom x y → Node x y - no : ∀ {x y} → Node' x y → Node x y - - data Node' where - _∷N_ : ∀ {x y z} → {∷Guard : N 𝒄𝒐𝒏𝒔} → Node x y → Atom y z → Node' x z - _⋆N_ : ∀ {x y z} → {⋆Guard : N 𝒐𝒑} → Node x y → Node y z → Node' x z - invN : ∀ {x y} → {invGuard : N 𝒊𝒏𝒗} → Node y x → {Bool→Type hasInvs} → Node' x y - involN : ∀ {x y z} → {g : N 𝒊𝒏𝒗𝒐𝒍} → Node x y → Atom y z → {Bool→Type hasInvs} → Node' x y - - - _a⤋_ : ∀ k → ∀ {x y} → Atom k x y → Atom (suc k) x y - k a⤋ a⟦ x ⟧ = a⟦ x ⟧ - k a⤋ (a⟦ x ⟧⁻ {g}) = a⟦_⟧⁻ {_} {_} {_} {g} x {g} - - - _N2++_ : ∀ {x y z k} → Node (suc (suc k)) x y - → Node (suc (suc k)) y z - → Node (suc (suc k)) x z - x N2++ idN = x - x N2++ atomN x₁ = no (x ∷N x₁) - x N2++ no (x₁ ∷N x₂) = no ((x N2++ x₁) ∷N x₂) - x N2++ no (involN {g = g} x₁ x₂ {uuu}) = no (involN {g = g} (x N2++ x₁) x₂ {uuu}) - - - invAtom : ∀ {x y} k {k'} → (Bool→Type hasInvs) → Atom k x y → Atom (suc k') y x - invAtom k x a⟦ x₁ ⟧ = a⟦_⟧⁻ {_} {_} {_} {x} x₁ {x} - invAtom k x a⟦ x₁ ⟧⁻ = a⟦ x₁ ⟧ - - invNode : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node k x y → Node k y x - invNode' : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node' k x y → Node k y x - - invNode' {k = k} {x} (_⋆N_ {⋆Guard = ⋆Guard} x₁ x₂) = - no (_⋆N_ {⋆Guard = ⋆Guard} (invNode {hi = x} x₂) (invNode {hi = x} x₁)) - invNode' {k = k} {x} (invN {invGuard = invGuard} x₁) = x₁ - invNode' {k = suc (suc k)} {x} (involN {g = gg} x₁ x₂ {g}) = - no (involN {g = gg} idN (x₂) {g}) N2++ invNode {hi = x} x₁ - invNode' {k = suc (suc k)} {x} (x₁ ∷N x₂) = - no (idN ∷N invAtom _ x x₂) N2++ invNode {hi = x} x₁ - invNode {k = k} {x} idN = idN - invNode {k = zero} {x} (atomN {aGuard = g} x₁) = - no (invN {invGuard = x} (atomN {aGuard = g} x₁) {x}) - invNode {k = suc k} {x} (atomN {aGuard = g} x₁) = - atomN {aGuard = g} (invAtom _ x x₁) - invNode {k = k} {x} (no x₁) = invNode' {hi = x} x₁ - - - invNode* : ∀ {x y k} → {hi : Bool→Type hasInvs} → Node k x y → Node k y x - invNode* {k = zero} {hi} f = no (invN {invGuard = hi} f {hi}) - invNode* {k = suc k} {hi} = invNode {k = suc k} {hi} - - - len : ∀ {x y} → Node 2 x y → ℕ - len idN = 0 - len (no (x ∷N x₁)) = suc (len x) - len (no (involN x x₁)) = suc (suc (len x)) - - module _ {k} (showH : {a₀ a₁ : ob} → (p : Hom[ a₀ , a₁ ]) → String) where - - showA showA' : {a₀ a₁ : ob} → Atom k a₀ a₁ → String - showA a⟦ x ⟧ = showH x - showA a⟦ x ⟧⁻ = "(" <> showH x <> ")⁻¹" - showA' a⟦ x ⟧ = "(" <> showH x <> ")⁻¹" - showA' a⟦ x ⟧⁻ = showH x - - showN : {a₀ a₁ : ob} → Node k a₀ a₁ → String - showN' : {a₀ a₁ : ob} → Node' k a₀ a₁ → String - - showN idN = "id" - showN (atomN x) = showA x - showN (no x) = showN' x - - showN' (x ∷N x₁) = showN x <> "⋆" <> showA x₁ - showN' (x ⋆N x₁) = "(" <> showN x <> "⋆" <> showN x₁ <> ")" - showN' (invN x) = "(" <> showN x <> ")⁻¹" - showN' (involN x x₁ {v}) = showN x <> - "⋆⟦" <> showA x₁ <> "⋆" <> showA' x₁ <> "⟧" - - - _⤋'_ : ∀ k → ∀ {x y} → Node' k x y → Node (suc k) x y - - _⤋_ : ∀ k → ∀ {x y} → Node k x y → Node (suc k) x y - k ⤋ idN = idN - zero ⤋ atomN x = atomN (0 a⤋ x) - suc zero ⤋ atomN x = no (idN ∷N (_ a⤋ x)) - k ⤋ no x = (k ⤋' x) - - suc (suc k) ⤋' (x ∷N x₁) = no ((suc (suc k) ⤋ x) ∷N (suc (suc k) a⤋ x₁)) - zero ⤋' (x ⋆N x₁) = no ((0 ⤋ x) ⋆N (0 ⤋ x₁)) - suc zero ⤋' (x ⋆N x₁) = - (suc zero ⤋ x) N2++ (suc zero ⤋ x₁) - zero ⤋' invN x {hi} = invNode {hi = hi} (zero ⤋ x) - - suc (suc k) ⤋' involN {_} {_} {_} {zz} x x₁ {zzz} = - no (involN {_} {_} {_} {_} {zz} (suc (suc k) ⤋ x) (suc (suc k) a⤋ x₁) {zzz}) - - _⤋⁺_ : ∀ {k} m → ∀ {x y} → Node k x y → Node (m + k) x y - zero ⤋⁺ x = x - suc m ⤋⁺ x = (m + _) ⤋ (m ⤋⁺ x) - - red[_,_] : ∀ k → ∀ {x y} → Node k x y → Node k x y - red[ zero , f ] = f - red[ suc zero , f ] = f - red[ suc (suc k) , idN ] = idN - red[ suc (suc k) , no (x ∷N x₁) ] = - no (red[ suc (suc k) , x ] ∷N x₁) - red[ suc (suc k) , no (involN x x₁) ] = red[ suc (suc k) , x ] - -module _ (A : Type ℓ) where - module Expr = Nodes Unit (λ _ _ → A) - - module DecNodes (_≟A_ : Discrete A) where - - AtomTo𝟚×A : Expr.Atom true 2 _ _ → (Bool × A) - AtomTo𝟚×A Nodes.a⟦ x ⟧ = true , x - AtomTo𝟚×A Nodes.a⟦ x ⟧⁻ = false , x - - - mbRed : Expr.Node true 2 _ _ → Maybe (Expr.Node true 2 _ _) - mbRed Nodes.idN = nothing - mbRed (Nodes.no (Nodes.idN Nodes.∷N x₁)) = nothing - mbRed (Nodes.no (x'@(Nodes.no (x Nodes.∷N x₂)) Nodes.∷N x₁)) = - decRec (λ _ → just $ Nodes.no (Nodes.involN (Mb.rec x (idfun _) (mbRed x)) x₂) ) - (λ _ → map-Maybe (Nodes.no ∘ Nodes._∷N x₁) (mbRed x')) - (discreteΣ 𝟚._≟_ (λ _ → _≟A_) (AtomTo𝟚×A x₂) (AtomTo𝟚×A (Expr.invAtom true 2 _ x₁))) - mbRed (Nodes.no (Nodes.no (Nodes.involN x x₂) Nodes.∷N x₁)) = nothing - mbRed (Nodes.no (Nodes.involN x x₁)) = nothing - - redList : Expr.Node true 2 _ _ → List (Expr.Node true 2 _ _) - redList x = unfoldMaybe (Expr.len true x) (mbRed ∘ Expr.red[_,_] true 2) x - -redListℕ = DecNodes.redList ℕ discreteℕ - -mapExpr : ∀ {A : Type ℓ} {A' : Type ℓ'} {b k} - → (A → A') - → Expr.Node A b k _ _ - → Expr.Node A' b k _ _ -mapExpr {A = A} {A'} {b} {k} f = w - where - wa : Expr.Atom A b k _ _ → Expr.Atom A' b k _ _ - wa Nodes.a⟦ x ⟧ = Nodes.a⟦ f x ⟧ - wa (Nodes.a⟦_⟧⁻ {invG = ig} x {g}) = Nodes.a⟦_⟧⁻ {invG = ig} (f x) {g} - - w : Expr.Node A b k _ _ → Expr.Node A' b k _ _ - w Nodes.idN = Nodes.idN - w (Nodes.atomN {aGuard = ag} x) = Nodes.atomN {aGuard = ag} (wa x) - w (Nodes.no (Nodes._∷N_ {∷Guard = g} x x₁)) = - Nodes.no (Nodes._∷N_ {∷Guard = g} (w x) (wa x₁)) - w (Nodes.no (Nodes._⋆N_ {⋆Guard = g} x x₁)) = - Nodes.no (Nodes._⋆N_ {⋆Guard = g} (w x) (w x₁)) - w (Nodes.no (Nodes.invN {invGuard = invGuard} x {g})) = - (Nodes.no (Nodes.invN {invGuard = invGuard} (w x) {g})) - w (Nodes.no (Nodes.involN {g = g} x x₁ {g'})) = - Nodes.no (Nodes.involN {g = g} (w x) (wa x₁) {g'}) - - -mapExprQ : ∀ {A : Type ℓ} {b k} - → (A → R.Term) → Expr.Node A b k _ _ → R.Term -mapExprQ {A = A} {b} {k} f = w - where - wa : Expr.Atom A b k _ _ → R.Term - wa Nodes.a⟦ x ⟧ = R.con (quote Nodes.a⟦_⟧) (f x v∷ []) - wa Nodes.a⟦ x ⟧⁻ = R.con (quote Nodes.a⟦_⟧⁻) (f x v∷ []) - - w : Expr.Node A b k _ _ → R.Term - w' : Expr.Node' A b k _ _ → R.Term - w' (x Nodes.∷N x₁) = - R.con (quote Nodes._∷N_) (w x v∷ wa x₁ v∷ []) - w' (x Nodes.⋆N x₁) = - R.con (quote Nodes._⋆N_) (w x v∷ w x₁ v∷ []) - w' (Nodes.invN x) = - R.con (quote Nodes.invN) (w x v∷ []) - w' (Nodes.involN x x₁) = - R.con (quote Nodes.involN) (w x v∷ wa x₁ v∷ []) - - w Nodes.idN = R.con (quote Nodes.idN) [] - w (Nodes.atomN x) = R.con (quote Nodes.atomN) (wa x v∷ []) - w (Nodes.no x) = R.con (quote Nodes.no) (w' x v∷ []) - - -ExprAccumM : ∀ {A : Type ℓ} {A' : Type ℓ'} {ℓs} {S : Type ℓs} {b k} - → (S → A → R.TC (S × A')) → S - → Expr.Node A b k _ _ - → R.TC (S × Expr.Node A' b k _ _) -ExprAccumM {A = A} {A'} {S = S} {b} {k} f = w - where - open Expr - - wa : S → Expr.Atom A b k _ _ → R.TC (S × Expr.Atom A' b k _ _) - wa s a⟦ x ⟧ = (λ (s' , x') → s' , a⟦ x' ⟧) <$> f s x - wa s (a⟦_⟧⁻ {invG = g'} x {g}) = (λ (s' , x') → s' , (a⟦_⟧⁻ {invG = g'} x' {g})) <$> f s x - - w : S → Expr.Node A b k _ _ → R.TC (S × Expr.Node A' b k _ _) - w' : S → Expr.Node' A b k _ _ → R.TC (S × Expr.Node A' b k _ _) - w s idN = R.returnTC (s , idN) - w s (atomN {aGuard = ag} a) = - (λ (s' , a') → s' , atomN {aGuard = ag} a') <$> wa s a - w s (no x) = w' s x - - w' s (_∷N_ {∷Guard = g} x x₁) = do - (s' , x') ← w s x - (s' , x₁') ← wa s' x₁ - R.returnTC (s' , no (_∷N_ {∷Guard = g} x' x₁')) - - w' s (_⋆N_ {⋆Guard = g} x x₁) = do - (s' , x') ← w s x - (s' , x₁') ← w s' x₁ - R.returnTC (s' , no (_⋆N_ {⋆Guard = g} x' x₁')) - - w' s (invN {invGuard = g'} x {g}) = - (λ (s' , x') → s' , (no (invN {invGuard = g'} x' {g}))) <$> w s x - w' s (involN x x₁) = w s x - - - -WildStr : ∀ ℓ ℓ' → Type (ℓ-suc (ℓ-max ℓ ℓ')) -WildStr ℓ ℓ' = Σ (WildCat ℓ ℓ') (Maybe ∘ IsWildGroupoid) - -hasInvs? : WildStr ℓ ℓ' → Bool -hasInvs? = caseMaybe false true ∘ snd - -WildStr→WG : (ws : WildStr ℓ ℓ') → Bool→Type (hasInvs? ws) → WildGroupoid ℓ ℓ' -WildGroupoid.wildCat (WildStr→WG (wc , _) _) = wc -WildGroupoid.isWildGroupoid (WildStr→WG (_ , just x) _) = x - - -module WSExpr (WS : WildStr ℓ ℓ') where - open WildCat (fst WS) public - open Nodes ob Hom[_,_] (hasInvs? WS) public - - InvGuard = Bool→Type (hasInvs? WS) - - module _ {ig : InvGuard} where - open WildGroupoid (WildStr→WG WS ig) public - - eva[_] : ∀ {k} → ∀ {x y} → Atom k x y → Hom[ x , y ] - eva[ a⟦ x ⟧ ] = x - eva[ a⟦ x ⟧⁻ {ig} ] = inv {ig} x - - ev[_] : ∀ {k} → ∀ {x y} → Node k x y → Hom[ x , y ] - ev[_]b : ∀ {k} → ∀ {x y} → Node k x y → Bool → Hom[ x , y ] - - ev[ f ] = ev[ f ]b false - - - ev[ idN ]b _ = id - ev[ atomN x ]b _ = eva[ x ] - ev[ no (xs ∷N x) ]b b = ev[ xs ]b b ⋆ eva[ x ] - ev[ no (x ⋆N x₁) ]b b = ev[ x ]b b ⋆ ev[ x₁ ]b b - ev[ no (invN x {invGuard}) ]b b = inv {invGuard} (ev[ x ]b b) - ev[ no (involN x x₁ {invGuard}) ]b false = (ev[ x ] ⋆ eva[ x₁ ]) ⋆ - eva[ invAtom _ {1} invGuard x₁ ] - ev[ no (involN x x₁) ]b true = ev[ x ]b true - - evInv : ∀ {k} → ∀ {x y hi} → (f : Atom k x y) → - id ≡ (eva[ f ] ⋆ eva[ invAtom k {1} hi f ]) - evInv a⟦ x ⟧ = sym (⋆InvR x) - evInv a⟦ x ⟧⁻ = sym (⋆InvL x) - - ev[_]≡ : ∀ {k} → ∀ {x y} → (f : Node k x y) → ev[ f ]b false ≡ ev[ f ]b true - ev[ idN ]≡ = refl - ev[ atomN x ]≡ = refl - ev[ no (x ∷N x₁) ]≡ = cong (_⋆ eva[ x₁ ]) ev[ x ]≡ - ev[ no (x ⋆N x₁) ]≡ = cong₂ _⋆_ ev[ x ]≡ ev[ x₁ ]≡ - ev[ no (invN x) ]≡ = cong inv ev[ x ]≡ - ev[ no (involN x x₁) ]≡ = - ⋆Assoc _ _ _ - ∙∙ cong₂ _⋆_ ev[ x ]≡ (sym (evInv x₁)) - ∙∙ ⋆IdR _ - - ev[_]₂≡ = ev[_]≡ {k = 2} - - - evN2++b : ∀ {k} {a₀ a₁ a₂ : ob} b → (x : Node (suc (suc k)) a₀ a₁) - → (x₁ : Node (suc (suc k)) a₁ a₂) - → ev[ x ]b b ⋆ ev[ x₁ ]b b ≡ ev[ x N2++ x₁ ]b b - evN2++b b x idN = ⋆IdR _ - evN2++b b x (no (x₁ ∷N x₂)) = - sym (⋆Assoc _ _ _) - ∙ cong (_⋆ eva[ x₂ ]) (evN2++b b x x₁) - evN2++b false x (no (involN x₁ x₂)) = - sym (⋆Assoc _ _ _) - ∙∙ congS (_⋆ eva[ invAtom _ _ x₂ ]) (sym (⋆Assoc _ _ _)) - ∙∙ congS (λ y → ((y ⋆ eva[ x₂ ]) ⋆ eva[ invAtom _ _ x₂ ])) (evN2++b false x x₁) - - evN2++b true x (no (involN x₁ x₂)) = evN2++b true x x₁ - - evN2++ : ∀ {k} {a₀ a₁ a₂ : ob} → (x : Node (suc (suc k)) a₀ a₁) - → (x₁ : Node (suc (suc k)) a₁ a₂) - → ev[ x ] ⋆ ev[ x₁ ] ≡ ev[ x N2++ x₁ ] - evN2++ = evN2++b false - - inv-eva : ∀ {x y} {k} {k'} a {hi} → - inv {hi} {x} {y} (eva[_] {k} a) ≡ eva[ invAtom (k) {k'} hi a ] - inv-eva a⟦ x ⟧ = refl - inv-eva (a⟦ x ⟧⁻ {hi}) {hi'} = - cong (inv {hi'}) ((λ i → inv {isPropBool→Type hi hi' i}) ≡$ x) ∙ invol-inv x - - invAtomInvol : ∀ {x y} k {hi hi'} → (x : Atom (suc (suc k)) x y) → - (eva[ invAtom 2 {suc k} (hi) (invAtom (suc (suc k)) {1} hi' x) ]) ≡ (eva[ x ]) - invAtomInvol k a⟦ x ⟧ = refl - invAtomInvol k {hi} (a⟦ x ⟧⁻ {hi'}) i = inv {isPropBool→Type hi hi' i} x - - - ev[_]≡inv : ∀ {k} → ∀ {x y} {hi} b → (f : Node k x y) → - inv {hi} (ev[ f ]b b) ≡ ev[ invNode* {hi = hi} f ]b b - ev[_]≡inv {zero} b f = refl - ev[_]≡inv {suc k} b idN = id≡inv-id - ev[_]≡inv {suc zero} {hi = hi} b (atomN x) = inv-eva x - ev[_]≡inv {suc zero} b (no (x ⋆N x₁)) = - distInv (ev[ x ]b b) (ev[ x₁ ]b b) ∙ - cong₂ _⋆_ (ev[_]≡inv b x₁) (ev[_]≡inv b x) - ev[_]≡inv {suc (suc k)} {hi = hi} b (no (x ∷N x₁)) = - distInv (ev[ x ]b b) (eva[ x₁ ]) - ∙∙ cong₂ _⋆_ (λ i → ⋆IdL (inv-eva {k' = suc k} x₁ {hi} i) (~ i)) - (ev[_]≡inv b x) - ∙∙ evN2++b b (no (idN ∷N invAtom (suc (suc k)) hi x₁)) (invNode x) - - ev[_]≡inv {suc (suc k)} {hi = hi} false (no (involN x x₁ {hi'})) = - distInv (ev[ x ] ⋆ eva[ x₁ ]) (eva[ invAtom (suc (suc k)) hi' x₁ ]) - ∙∙ - cong₂ _⋆_ - (inv-eva {k' = suc k} (invAtom (suc (suc k)) hi' x₁)) - (distInv ev[ x ] eva[ x₁ ]) ∙ - (λ i → ⋆Assoc (⋆IdL (invAtomInvol k {hi} {hi'} x₁ i) (~ i)) - (inv-eva {k' = 1} x₁ {isPropBool→Type hi hi' i} i) - (ev[_]≡inv {suc (suc k)} {hi = hi} false x i) (~ i) ) - ∙∙ evN2++b false (no (involN idN x₁)) (invNode x) - - ev[_]≡inv {suc (suc k)} true (no (involN x x₁)) = - ev[_]≡inv true x - ∙∙ sym (⋆IdL _) - ∙∙ evN2++b true (no (involN idN x₁)) (invNode x) - - - data NodeCase : {a₀ a₁ : ob} → Hom[ a₀ , a₁ ] → Type (ℓ-max ℓ ℓ') where - idCase : ∀ {x} → NodeCase (id {x = x}) - opCase : ∀ {x y z : _} → (p : Hom[ x , y ]) (q : Hom[ y , z ]) → NodeCase (p ⋆ q) - invCase : ∀ {x y : _} → {hi : InvGuard} (p : Hom[ x , y ]) → NodeCase (inv {hi} p) - - - eva⤋ : ∀ k {a₀ a₁ : ob} → ∀ (h : Atom k a₀ a₁) → eva[ h ] ≡ eva[ k a⤋ h ] - eva⤋ k a⟦ x ⟧ = refl - eva⤋ k a⟦ x ⟧⁻ = refl - - - invAtom⤋ : ∀ k k' {hi} {a₀ a₁ : ob} → ∀ (h : Atom (suc (suc k)) a₀ a₁) → - eva[ invAtom (suc (suc k)) {k'} hi h ] - ≡ eva[ invAtom (suc (suc (suc k))) {k'} hi (suc (suc k) a⤋ h) ] - invAtom⤋ k k' Nodes.a⟦ x ⟧ = refl - invAtom⤋ k k' Nodes.a⟦ x ⟧⁻ = refl - - - ev⤋' : ∀ k {a₀ a₁ : ob} → ∀ (f' : Node' k a₀ a₁) → ev[ no f' ] ≡ ev[ k ⤋' f' ] - ev⤋ : ∀ k {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ k ⤋ f ] - - - - ev⤋ k idN = refl - ev⤋ zero (Nodes.atomN Nodes.a⟦ x ⟧) = refl - ev⤋ (suc zero) (Nodes.atomN Nodes.a⟦ x ⟧) = sym (⋆IdL x) - ev⤋ (suc zero) (Nodes.atomN (Nodes.a⟦ x ⟧⁻)) = sym ((⋆IdL _)) - ev⤋ k (no x) = ev⤋' k x - - ev⤋' (suc (suc k)) (x ∷N x₁) = - cong₂ _⋆_ (ev⤋ (suc (suc k)) x) (eva⤋ (suc (suc k)) x₁ ) - ev⤋' zero (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ zero x) (ev⤋ zero x₁) - ev⤋' (suc zero) (x ⋆N x₁) = cong₂ _⋆_ (ev⤋ 1 x) (ev⤋ 1 x₁) ∙ - evN2++ (1 ⤋ x) (1 ⤋ x₁) - ev⤋' zero (invN x {hi}) = - cong inv (ev⤋ zero x) ∙ ev[_]≡inv false (0 ⤋ x) -- enInv1Node x hi - ev⤋' (suc (suc k)) (Nodes.involN x x₁ {hi}) = - cong₂ _⋆_ (λ i → (ev⤋ (suc (suc k)) x i ⋆ (eva⤋ (suc (suc k)) x₁) i)) - (invAtom⤋ k 1 {hi} x₁) - - ev⤋⁺ : ∀ {k} m {a₀ a₁ : ob} → ∀ (f : Node k a₀ a₁) → ev[ f ] ≡ ev[ m ⤋⁺ f ] - ev⤋⁺ zero f = refl - ev⤋⁺ (suc m) f = ev⤋⁺ m f ∙ ev⤋ (m + _) (m ⤋⁺ f) - - ev⤋² = ev⤋⁺ {0} 2 - -module Eℕ = Expr ℕ true -module ETerm = Expr R.Term - -NodeTerm : Bool → ℕ → Type ℓ-zero -NodeTerm = λ b k → Expr.Node R.Term b k tt tt - - - -module tryWCE (tG : R.Term) where - - tryG : ℕ → R.Term → R.TC (NodeTerm true 0) - - try1g : R.Term → R.TC (NodeTerm true 0) - try1g t = do - _ ← R.unify t (R.def (quote WSExpr.id) [ varg tG ]) - R.returnTC (ETerm.idN) - - tryOp : ℕ → R.Term → R.TC (NodeTerm true 0) - tryOp zero _ = R.typeError [] - tryOp (suc k) t = do - tm ← R.withNormalisation true $ R.checkType (R.con (quote WSExpr.opCase) - (varg R.unknown ∷ [ varg R.unknown ])) - (R.def (quote WSExpr.NodeCase) ((varg tG) ∷ [ varg t ])) - (t1 , t2) ← h tm - t1' ← tryG k t1 - t2' ← tryG k t2 - R.returnTC (ETerm.no (t1' ETerm.⋆N t2')) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con _ l) = match2Vargs l - h t = R.typeError [] - - tryInv : ℕ → R.Term → R.TC (NodeTerm true 0) - tryInv zero _ = R.typeError [] - tryInv (suc k) t = do - tm ← R.withNormalisation true $ - (R.checkType (R.con (quote WSExpr.invCase) - ([ varg R.unknown ])) (R.def (quote WSExpr.NodeCase) - ((varg tG) ∷ [ varg t ]))) - R.debugPrint "tryInv" 30 ([ R.strErr "\n ---- \n" ]) - R.debugPrint "tryInv" 30 ([ R.termErr t ]) - R.debugPrint "tryInv" 30 ([ R.termErr tm ]) - t1 ← h tm - t1' ← tryG k t1 - R.returnTC (ETerm.no (ETerm.invN t1')) - - where - h' : List (R.Arg R.Term) → R.TC (R.Term) - h' (harg _ ∷ xs) = h' xs - h' (varg t1 ∷ []) = R.returnTC t1 - h' _ = do R.debugPrint "aiu" 3 ([ R.strErr "\n vvv \n" ]) - R.debugPrint "aiu" 33 ([ R.termErr t ]) - R.typeError [ R.strErr "yyy" ] - - h : R.Term → R.TC (R.Term) - h (R.con _ l) = h' l - h t = do R.debugPrint "aiu" 33 ([ R.strErr "\n uuu \n" ]) - R.debugPrint "aiu" 33 ([ R.termErr t ]) - R.typeError [ R.strErr "xxxx" ] - - - atom : R.Term → R.TC (NodeTerm true 0) - atom x = R.returnTC $ ETerm.atomN ETerm.a⟦ x ⟧ - - - tryG zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryG (suc k) t = - R.catchTC - (try1g t) - (R.catchTC (tryInv k t) - -- (tryOp k t) - (R.catchTC (tryOp k t) (atom t)) - ) - - - - - -compareTerms : R.Term → R.Term → R.TC Bool -compareTerms x x' = - ((R.noConstraints $ R.unify x x') >> (R.returnTC true)) <|> - (R.returnTC false) - - - -lookupOrAppend : List R.Term → R.Term → R.TC ((List R.Term) × ℕ) -lookupOrAppend [] t = R.returnTC ([ t ] , 0) -lookupOrAppend xs@(x ∷ xs') t = do - x≟t ← compareTerms t x - if x≟t - then (R.returnTC (xs , 0)) - else (do (xs'' , k) ← lookupOrAppend xs' t - R.returnTC (x ∷ xs'' , suc k)) - - - -wildCatSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) -wildCatSolverTerm debugFlag t-g hole = do - - (t0 , t1) ← R.withNormalisation true $ - R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - (tmL , tmE0) ← tryWCE.tryG t-g 100 t0 >>= ExprAccumM lookupOrAppend [] - (tmL , tmE1) ← tryWCE.tryG t-g 100 t1 >>= ExprAccumM lookupOrAppend tmL - - let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) - pa = λ tmE → - let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) - rLpaTm = foldl - (λ x y → - (R∙ x ( R.def (quote WSExpr.ev[_]₂≡) - (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) - Rrefl rL - in ((R.def (quote WSExpr.ev⤋²) - (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ [])) , - (Li.foldr - (λ x → Expr.showN ℕ true mkNiceVar x ∷nl_ ) [] $ rL)) - - let final = (R.def (quote _∙∙_∙∙_) - (fst (pa tmE0) - v∷ R.def (quote refl) [] - v∷ R.def (quote sym) (fst (pa tmE1) v∷ []) v∷ [] )) - - - let info = snd (pa tmE0) ++ snd (pa tmE1) ++ (Expr.showN ℕ true mkNiceVar tmE0 - ∷nl Expr.showN ℕ true mkNiceVar tmE1 - ∷nl niceAtomList tmL) - - R.returnTC (final , info) - - -wildCatSolverMain : Bool → R.Term → R.Term → R.TC Unit -wildCatSolverMain debugFlag t-g hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← wildCatSolverTerm debugFlag t-g hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - - - - -macro - solveWildCat : R.Term → R.Term → R.TC Unit - solveWildCat = wildCatSolverMain true - - - -groupoidSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag t-g hole = do - - (t0 , t1) ← R.withNormalisation true $ - R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - (tmL , tmE0) ← tryWCE.tryG t-g 100 t0 >>= ExprAccumM lookupOrAppend [] - (tmL , tmE1) ← tryWCE.tryG t-g 100 t1 >>= ExprAccumM lookupOrAppend tmL - - let pa : Eℕ.Node 0 _ _ → (R.Term × List R.ErrorPart) - pa = λ tmE → - let rL = redListℕ (1 Eℕ.⤋ (0 Eℕ.⤋ tmE)) - rLpaTm = foldl - (λ x y → - (R∙ x ( R.def (quote WSExpr.ev[_]₂≡) - (t-g v∷ (mapExprQ (lookupWithDefault (R.unknown) tmL) y) v∷ []))) ) - Rrefl rL - in ((R.def (quote _∙_) - (R.def (quote WSExpr.ev⤋²) - (t-g v∷ mapExprQ (lookupWithDefault (R.unknown) tmL) tmE v∷ []) - v∷ rLpaTm v∷ [] )) , - (Li.foldr - (λ x → Expr.showN ℕ true mkNiceVar x ∷nl_ ) [] $ rL)) - - let final = (R.def (quote _∙∙_∙∙_) - (fst (pa tmE0) - v∷ R.def (quote refl) [] - v∷ R.def (quote sym) (fst (pa tmE1) v∷ []) v∷ [] )) - - - let info = snd (pa tmE0) ++ snd (pa tmE1) ++ (Expr.showN ℕ true mkNiceVar tmE0 - ∷nl Expr.showN ℕ true mkNiceVar tmE1 - ∷nl niceAtomList tmL) - - R.returnTC (final , info) - - -groupoidSolverMain : Bool → R.Term → R.Term → R.TC Unit -groupoidSolverMain debugFlag t-g hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag t-g hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - - - - -macro - solveGroupoid : R.Term → R.Term → R.TC Unit - solveGroupoid = groupoidSolverMain true - - - - - - - - - - - - - - - diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda new file mode 100644 index 0000000000..8000876cfa --- /dev/null +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -0,0 +1,170 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.WildCatSolver.Example where + +open import Cubical.Foundations.Prelude + +open import Cubical.WildCat.Base +open import Cubical.Tactics.WildCatSolver.Solver +open import Cubical.Data.List + +open import Cubical.WildCat.Functor + +private + variable + ℓ ℓ' : Level + +module exampleC (WC WC* : WildCat ℓ ℓ') + (F : WildFunctor WC* WC) where + + open WildCat WC + module * = WildCat WC* + + module _ x y z w v + (p : Hom[ x , F ⟅ y ⟆ ]) + (q : *.Hom[ y , z ]) + (r : *.Hom[ z , w ]) + (s : Hom[ F ⟅ w ⟆ , v ]) where + + + pA pB pC : Hom[ x , v ] + pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) + pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) + pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) + + open WildCat-Solver ℓ ℓ' + + pA=pB : pA ≡ pB + pA=pB = solveWildCat (WC ∷ WC* ∷ []) + + pB=pC : pB ≡ pC + pB=pC = solveWildCat (WC ∷ WC* ∷ []) + + + +-- -- module example (WG WG* : WildGroupoid ℓ ℓ') +-- -- (F : WildFunctor +-- -- (WildGroupoid.wildCat WG*) +-- -- (WildGroupoid.wildCat WG)) +-- -- where + +-- -- module WildStrGPD = WildStr (GroupoidWS {ℓ} {ℓ'}) + +-- -- -- open WildStrGPD using +-- -- -- (testInferExpr; solveNoInvs; solveW) + +-- -- open WildGroupoid WG +-- -- open WildCat wildCat + +-- -- open WildFunctor + + +-- -- module * where +-- -- open WildCat (WildGroupoid.wildCat WG*) public +-- -- open WildGroupoid WG* public + +-- -- -- module T1 x y z w +-- -- -- (p : Hom[ x , y ]) +-- -- -- (q : Hom[ y , z ]) +-- -- -- (r : Hom[ z , w ]) where + + +-- -- -- pA pB : Hom[ x , w ] +-- -- -- pA = p ⋆ (q ⋆ r) +-- -- -- pB = (p ⋆ q) ⋆ r + +-- -- -- -- pF : + +-- -- -- pXX : Hom[ x , x ] +-- -- -- pXX = id + + +-- -- -- tryTestOp : Unit +-- -- -- tryTestOp = testTryOp GroupoidWS WG pA + +-- -- -- tryTestId : Unit +-- -- -- tryTestId = testTryId GroupoidWS WG pXX + +-- -- -- pACase : FuCases WG pA +-- -- -- pACase = p FuCases.⋆FE (q ⋆ r) + + +-- -- module T1 x y z w +-- -- (p p' : Hom[ x , y ]) +-- -- (q : Hom[ y , z ]) +-- -- (r : Hom[ z , w ]) where + + +-- -- -- pA pB : Hom[ x , x ] +-- -- -- pA = (p ⋆ inv p) +-- -- -- pB = id + +-- -- -- pA≡pB : pA ≡ pB +-- -- -- pA≡pB = {!solveW GroupoidWS (WG ∷ WG* ∷ [])!} + + +-- -- pA pB : Hom[ x , w ] +-- -- pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r +-- -- -- pA = ((((((id ⋆ p))) ⋆ (inv p)) ⋆ p) ⋆ q) ⋆ r +-- -- pB = ((id ⋆ p) ⋆ q) ⋆ r + +-- -- pA≡pB : pA ≡ pB +-- -- pA≡pB = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) + + + +-- -- module T2 x y z w +-- -- (p : Hom[ x , F ⟅ y ⟆ ]) +-- -- (p' : Hom[ F ⟅ y ⟆ , x ]) +-- -- (q : *.Hom[ z , y ]) +-- -- (r : *.Hom[ z , w ]) where + + +-- -- pA : Hom[ F ⟅ y ⟆ , F-ob F w ] +-- -- pA = F-hom F ((*.inv q) *.⋆ r) +-- -- pB pB' pB'' : Hom[ x , F ⟅ w ⟆ ] +-- -- pB = (p ⋆ (F-hom F ((*.inv q) *.⋆ r))) +-- -- pB'' = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) +-- -- pB' = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) + + +-- -- pB''≡pB' : pB'' ≡ pB' +-- -- pB''≡pB' = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) + + + + + + +-- -- -- module example (WG WG* : WildGroupoid ℓ ℓ') +-- -- -- (F : WildFunctor (WildGroupoid.wildCat WG*) (WildGroupoid.wildCat WG)) where +-- -- -- open WildGroupoid WG +-- -- -- open WildCat wildCat + + + +-- -- -- module * where +-- -- -- open WildCat (WildGroupoid.wildCat WG*) public +-- -- -- open WildGroupoid WG* public + +-- -- -- module _ x v y z w +-- -- -- (p p' : Hom[ x , F ⟅ y ⟆ ]) +-- -- -- (q q' : *.Hom[ y , z ]) +-- -- -- (r : *.Hom[ w , z ]) +-- -- -- (s : Hom[ F ⟅ w ⟆ , v ]) where + + +-- -- -- pA pB pC : Hom[ x , v ] +-- -- -- pA = (p ⋆ (id ⋆ F ⟪ q *.⋆ *.inv q' ⟫)) ⋆ (F ⟪ q' *.⋆ (*.inv r) ⟫ ⋆ s) +-- -- -- pB = ((p ⋆ F ⟪ q *.⋆ (*.inv (*.inv r *.⋆ r) *.⋆ *.inv q)⟫) +-- -- -- ⋆ F ⟪ q *.⋆ *.id ⟫) ⋆ ((inv (F ⟪ r ⟫ ⋆ id)) ⋆ s) + +-- -- -- pC = p' ⋆ (inv p' ⋆ (p ⋆ (F ⟪ q *.⋆ (*.inv r) ⟫ ⋆ s))) + +-- -- -- pA=pB : pA ≡ pB +-- -- -- pA=pB = solveGroupoid WG WG* + +-- -- -- pA=pC : pA ≡ pC +-- -- -- pA=pC = solveGroupoid WG WG* + + diff --git a/Cubical/Tactics/GroupoidSolver/Reflection.agda b/Cubical/Tactics/WildCatSolver/Reflection.agda similarity index 77% rename from Cubical/Tactics/GroupoidSolver/Reflection.agda rename to Cubical/Tactics/WildCatSolver/Reflection.agda index cabd7aaf38..f0b2c24468 100644 --- a/Cubical/Tactics/GroupoidSolver/Reflection.agda +++ b/Cubical/Tactics/WildCatSolver/Reflection.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Tactics.GroupoidSolver.Reflection where +module Cubical.Tactics.WildCatSolver.Reflection where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -13,7 +13,7 @@ open import Cubical.Data.Nat open import Cubical.Data.Unit open import Cubical.Data.Sigma open import Cubical.Data.List -open import Cubical.Data.Maybe +open import Cubical.Data.Maybe as Mb open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R @@ -124,6 +124,11 @@ match2Vargs (harg _ ∷ xs) = match2Vargs xs match2Vargs (varg t1 ∷ varg t2 ∷ []) = R.returnTC (t1 , t2) match2Vargs _ = R.typeError [] +matchFunctorAppArgs : List (R.Arg R.Term) → Maybe (R.Term × R.Term) +matchFunctorAppArgs (harg _ ∷ xs) = matchFunctorAppArgs xs +matchFunctorAppArgs (varg t1 ∷ harg _ ∷ harg _ ∷ varg t2 ∷ []) = just (t1 , t2) +matchFunctorAppArgs _ = nothing + match3Vargs : List (R.Arg R.Term) → R.TC (R.Term × R.Term × R.Term) match3Vargs (harg _ ∷ xs) = match3Vargs xs @@ -156,27 +161,43 @@ mkNiceVar : ℕ → String mkNiceVar k = "𝒙" <> primStringFromList (map digitsToSubscripts $ primStringToList $ primShowNat k) +mkNiceVar' : String → ℕ → String +mkNiceVar' v k = v <> + primStringFromList (map digitsToSubscripts $ primStringToList $ primShowNat k) + + record ToErrorPart {ℓ} (A : Type ℓ) : Type ℓ where field toErrorPart : A → R.ErrorPart open ToErrorPart -infixr 5 _∷ₑ_ _∷nl_ +infixr 5 _∷ₑ_ _∷nl_ _++ₑ_ + _∷ₑ_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart _∷ₑ_ ⦃ tep ⦄ x = (toErrorPart tep x) ∷_ +_++ₑ_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → List A → List R.ErrorPart → List R.ErrorPart +_++ₑ_ ⦃ tep ⦄ x = (map (toErrorPart tep) x) ++_ + instance ToErrorPartString : ToErrorPart String ToErrorPart.toErrorPart ToErrorPartString = R.strErr + ToErrorPartℕ : ToErrorPart ℕ + ToErrorPart.toErrorPart ToErrorPartℕ = R.strErr ∘ primShowNat + + ToErrorPartTerm : ToErrorPart R.Term ToErrorPart.toErrorPart ToErrorPartTerm = R.termErr ToErrorPartName : ToErrorPart R.Name ToErrorPart.toErrorPart ToErrorPartName = R.nameErr + ToErrorPartErrorPart : ToErrorPart R.ErrorPart + ToErrorPart.toErrorPart ToErrorPartErrorPart x = x + _∷nl_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart _∷nl_ x y = x ∷ₑ "\n" ∷ₑ y @@ -200,3 +221,29 @@ R∙ x y = R.def (quote _∙_) (x v∷ y v∷ [] ) Rrefl : R.Term Rrefl = R.def (quote refl) [] + +unArgs : List (R.Arg (R.Term)) → List R.ErrorPart +unArgs [] = [] +unArgs (R.arg i x ∷ x₁) = x ∷ₑ unArgs x₁ + +getConTail : R.Term → List R.ErrorPart +getConTail (R.var x args) = "𝒗𝒂𝒓 " ∷ₑ x ∷ₑ " " ∷ₑ unArgs args +getConTail (R.con c args) = "𝒄𝒐𝒏 " ∷ₑ c ∷ₑ " " ∷ₑ unArgs args +getConTail (R.def f args) = "𝒅𝒆𝒇 " ∷ₑ f ∷ₑ " " ∷ₑ unArgs args +getConTail _ = "other..." ∷ₑ [] + +tryAllTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → + R.TC B → List A → (A → R.TC B) → R.TC B +tryAllTC fallback [] f = fallback +tryAllTC fallback (x ∷ xs) f = + f x <|> tryAllTC fallback xs f + + +foldPathTerms : List (Maybe R.Term) → Maybe R.Term +foldPathTerms [] = nothing +foldPathTerms (nothing ∷ xs) = foldPathTerms xs +foldPathTerms (just x ∷ xs) = + just $ Mb.rec x (λ xs' → R.def (quote _∙_) (x v∷ xs' v∷ [])) (foldPathTerms xs) + +symPathTerms : List (Maybe R.Term) → List (Maybe R.Term) +symPathTerms = map (map-Maybe (R.def (quote sym) ∘ v[_])) ∘ rev diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda new file mode 100644 index 0000000000..d92216c4c0 --- /dev/null +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -0,0 +1,57 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.WildCatSolver.Solver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function as Fu +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +-- open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.GroupoidSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + +-- open import Cubical.WildCat.WGE +open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor +open import Cubical.Algebra.Group + +open import Cubical.Tactics.WildCatSolver.Solvers + + +module WildCat-Solver ℓ ℓ' where + + WildCatWS : WildStr ℓ ℓ' + WildStr.wildStr (WildCatWS) = WildCat ℓ ℓ' + WildStr.toWildCat WildCatWS = idfun _ + WildStr.mbIsWildGroupoid WildCatWS = nothing + + private + module WC-WS = WildStr WildCatWS + + macro + solveWildCat : R.Term → R.Term → R.TC Unit + solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) + diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda new file mode 100644 index 0000000000..7af64e3e02 --- /dev/null +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -0,0 +1,833 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.WildCatSolver.Solvers where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function as Fu +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +-- open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.GroupoidSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + +-- open import Cubical.WildCat.WGE +open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor +open import Cubical.Algebra.Group + +private + variable + ℓ ℓ' : Level + +mvFlagElim : ∀ {A : Type ℓ} (mbA : Maybe A) + → (caseMaybe {A = A} ⊥ Unit mbA) → A +mvFlagElim (just x) _ = x + + +record TypeWithRel ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + no-eta-equality + field + Carrier : Type ℓ + _[_∼_] : Carrier → Carrier → Type ℓ' + -- mbIsWildGroupoid : Maybe (∀ WS → IsWildGroupoid (toWildCat WS)) + + +record TypeWithRelMor (T T' : TypeWithRel ℓ ℓ') + : Type (ℓ-max ℓ ℓ') where + open TypeWithRel + field + TF-ob : TypeWithRel.Carrier T → TypeWithRel.Carrier T' + TF-hom : ∀ {x y} → T [ x ∼ y ] + → T' [ TF-ob x ∼ TF-ob y ] + + +FreeTWRM : ∀ {ℓ} (A : Type ℓ) → TypeWithRel ℓ-zero ℓ +TypeWithRel.Carrier (FreeTWRM A) = Unit +FreeTWRM A TypeWithRel.[ x ∼ x₁ ] = A + + + + +module FuExpr' (ℓ ℓ' : Level) (InvFlag : Type) + (𝑻 : Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ'))) + (𝑻→twr : 𝑻 → TypeWithRel ℓ ℓ') + (𝑭 : 𝑻 → 𝑻 → Type (ℓ-max ℓ ℓ')) + (𝑭-ob-map : ∀ {T T'} → 𝑭 T T' + → TypeWithRel.Carrier (𝑻→twr T) + → TypeWithRel.Carrier (𝑻→twr T') ) where + + module _ 𝑻 where open TypeWithRel (𝑻→twr 𝑻) public + + + data FuExpr : (T : 𝑻) → Carrier T → Carrier T → Type (ℓ-suc (ℓ-max ℓ ℓ')) where + 𝒂⟦_⟧ : ∀ {T x y} → T [ x ∼ y ] → FuExpr T x y + idFE : ∀ {T x} → FuExpr T x x + _⋆FE_ : ∀ {T x y z} + → FuExpr T x y + → FuExpr T y z → FuExpr T x z + invFE : ∀ {T x y} (invF : InvFlag) + → FuExpr T x y + → FuExpr T y x + ⟅_,_,_⟆FE : ∀ T {T' x y} + (F : 𝑭 T T') + → FuExpr T x y + → FuExpr T' (𝑭-ob-map F x) (𝑭-ob-map F y) + + + module _ (m2str : ∀ {T x y} → T [ x ∼ y ] → String) where + printFuExpr : ∀ {T x y} → FuExpr T x y → String + printFuExpr 𝒂⟦ x ⟧ = m2str x + printFuExpr idFE = "id" + printFuExpr (x ⋆FE x₁) = + "(" <> printFuExpr x <> "⋆" <> printFuExpr x₁ <> ")" + printFuExpr (invFE invF x) = + "(" <> printFuExpr x <> ")⁻¹" + printFuExpr ⟅ T , F , x ⟆FE = + "⟪" <> printFuExpr x <> "⟫" + +module _ InvFlag where + module TermExpr = FuExpr' ℓ-zero ℓ-zero InvFlag + (Lift R.Term) (λ _ → FreeTWRM R.Term) + (λ _ _ → R.Term) (λ _ _ → tt) + open TermExpr using (𝒂⟦_⟧; idFE; _⋆FE_; invFE; ⟅_,_,_⟆FE) + renaming (FuExpr to TE) public + +module _ InvFlag where + module ℕExpr = FuExpr' ℓ-zero ℓ-zero InvFlag + (Lift ℕ) (λ _ → FreeTWRM ℕ) + (λ _ _ → ℕ) (λ _ _ → tt) + + open ℕExpr using () + renaming (FuExpr to ℕE) public + +module _ if ([w] [F] [a] : List R.Term) where + lkW lkF lkA : ℕ → R.Term + lkW = lookupWithDefault R.unknown [w] + lkF = lookupWithDefault R.unknown [F] + lkA = lookupWithDefault R.unknown [a] + + ℕExpr→TermExrp : ∀ {w} → ℕE if (lift w) _ _ → TE if (lift (lkW w)) _ _ + ℕExpr→TermExrp 𝒂⟦ x ⟧ = 𝒂⟦ lkA x ⟧ + ℕExpr→TermExrp idFE = idFE + ℕExpr→TermExrp (x ⋆FE x₁) = + (ℕExpr→TermExrp x ⋆FE ℕExpr→TermExrp x₁) + ℕExpr→TermExrp (invFE invF x) = + invFE invF (ℕExpr→TermExrp x) + ℕExpr→TermExrp (⟅ _ , F , x ⟆FE) = ⟅ _ , lkF F , ℕExpr→TermExrp x ⟆FE + + +wc→twr : WildCat ℓ ℓ' → TypeWithRel ℓ ℓ' +TypeWithRel.Carrier (wc→twr x) = WildCat.ob x +TypeWithRel._[_∼_] (wc→twr x) = WildCat.Hom[_,_] x + +mbFunctorApp : R.Term → Maybe (R.Term × R.Term) +mbFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t +mbFunctorApp _ = nothing + +record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where + no-eta-equality + field + wildStr : Type (ℓ-suc (ℓ-max ℓ ℓ')) + toWildCat : wildStr → WildCat ℓ ℓ' + mbIsWildGroupoid : Maybe (∀ WS → IsWildGroupoid (toWildCat WS)) + + InvFlag = caseMaybe ⊥ Unit mbIsWildGroupoid + + ws→twr : wildStr → TypeWithRel ℓ ℓ' + ws→twr = wc→twr ∘ toWildCat + + module _ (ws : wildStr) where + open WildCat (toWildCat ws) renaming (Hom[_,_] to _H[_,_] ; _⋆_ to _⟨_⋆_⟩) public + module _ (invF : InvFlag) where + WG = (record { wildCat = toWildCat ws + ; isWildGroupoid = mvFlagElim mbIsWildGroupoid invF ws }) + open WildGroupoid WG public + + + WildF : wildStr → wildStr → Type (ℓ-max ℓ ℓ') + WildF ws ws' = WildFunctor (toWildCat ws) (toWildCat ws') + + open WildFunctor + + + data FuCases : (ws : wildStr) {x y : ob ws} + → ws H[ x , y ] → Type (ℓ-suc (ℓ-max ℓ ℓ')) where + 𝒂⟦_⟧ : ∀ {ws x y} f → FuCases ws {x} {y} f + idFE : ∀ {ws x} → FuCases ws {x} {x} (id ws) + _⋆FE_ : ∀ {ws x y z} + → (f : ws H[ x , y ]) + → (g : ws H[ y , z ]) → FuCases ws {x} {z} (ws ⟨ f ⋆ g ⟩) + invFE : ∀ {ws x y} (invF : InvFlag) + → (f : ws H[ x , y ]) + → FuCases ws {y} {x} (inv ws invF f) + ⟅_,_,_⟆FE : ∀ ws {ws' x y} + (F : WildF ws ws') + → (f : ws H[ x , y ]) + → FuCases ws' {F-ob F x} {F-ob F y} (F-hom F f ) + + module _ where + open FuExpr' ℓ ℓ' InvFlag wildStr ws→twr + WildF WildFunctor.F-ob public + + evFuExpr : {ws : wildStr} {x y : ob ws} + → FuExpr ws x y → ws H[ x , y ] + evFuExpr FuExpr'.𝒂⟦ x ⟧ = x + evFuExpr {ws} FuExpr'.idFE = id ws + evFuExpr {ws} (x FuExpr'.⋆FE x₁) = ws ⟨ evFuExpr x ⋆ evFuExpr x₁ ⟩ + evFuExpr {ws} (FuExpr'.invFE invF x) = inv ws invF (evFuExpr x) + evFuExpr FuExpr'.⟅ T , F , x ⟆FE = F ⟪ evFuExpr x ⟫ + + module _ {ws : wildStr} where + + data FuAtom : ob ws → ob ws → Type (ℓ-max ℓ ℓ') where + a[_] : ∀ {x y} → ws H[ x , y ] → FuAtom x y + a[_,_]⁻ : ∀ {x y} → InvFlag → ws H[ y , x ] → FuAtom x y + + + infixl 15 _ff∷_ + data FuFlat : ob ws → ob ws → Type (ℓ-max ℓ ℓ') where + [ff] : ∀ {x} → FuFlat x x + _ff∷_ : ∀ {x y z} → FuFlat x y → FuAtom y z → FuFlat x z + _invol_∷ff_ : ∀ {x y z} → FuFlat x y → (invF : InvFlag) → + FuAtom y z → FuFlat x y + + + invFuAtom : ∀ {x y} → InvFlag → FuAtom y x → FuAtom x y + invFuAtom x a[ x₁ ] = a[ x , x₁ ]⁻ + invFuAtom x a[ x₁ , x₂ ]⁻ = a[ x₂ ] + + + ff↓ : ∀ {x y} → FuFlat x y → FuFlat x y + ff↓ [ff] = [ff] + ff↓ (x ff∷ x₁) = ff↓ x ff∷ x₁ + ff↓ (x invol invF ∷ff x₁) = ff↓ x ff∷ x₁ ff∷ invFuAtom invF x₁ + + ff↑ : ∀ {x y} → FuFlat x y → FuFlat x y + ff↑ [ff] = [ff] + ff↑ (x ff∷ x₁) = ff↑ x ff∷ x₁ + ff↑ (x invol invF ∷ff x₁) = ff↑ x + + + _ff++_ : ∀ {x y z} → FuFlat x y → FuFlat y z → FuFlat x z + x ff++ [ff] = x + x ff++ (x₁ ff∷ x₂) = (x ff++ x₁) ff∷ x₂ + x ff++ (x₁ invol invF ∷ff x₂) = (x ff++ x₁) invol invF ∷ff x₂ + + + ffInv : ∀ {x y} → InvFlag → FuFlat x y → FuFlat y x + ffInv x [ff] = [ff] + ffInv x (x₁ ff∷ x₂) = ([ff] ff∷ (invFuAtom x x₂)) ff++ ffInv x x₁ + ffInv x (x₁ invol invF ∷ff x₂) = + ([ff] invol invF ∷ff x₂) ff++ ffInv x x₁ + + invFuAtomExplicit : ∀ (ws : wildStr) {x y : WildCat.ob (toWildCat ws)} → + InvFlag → FuAtom y x → FuAtom x y + invFuAtomExplicit ws = invFuAtom {ws} + + aa⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} + → (F : WildF ws' ws) + → (FuAtom x y) + → FuAtom (F-ob F x) (F-ob F y) + aa⟪ F , a[ x ] ⟫ = a[ F ⟪ x ⟫ ] + aa⟪ F , a[ x , x₁ ]⁻ ⟫ = a[ x , F ⟪ x₁ ⟫ ]⁻ + + + ff⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} + → (F : WildF ws' ws) + → (FuFlat x y) + → FuFlat (F-ob F x) (F-ob F y) + ff⟪ F , [ff] ⟫ = [ff] + ff⟪ F , x ff∷ x₁ ⟫ = ff⟪ F , x ⟫ ff∷ aa⟪ F , x₁ ⟫ + ff⟪ F , x invol invF ∷ff x₁ ⟫ = ff⟪ F , x ⟫ invol invF ∷ff aa⟪ F , x₁ ⟫ + + + FuExpr→FF : {ws : wildStr} {x y : ob ws} + → FuExpr ws x y → FuFlat x y + FuExpr→FF 𝒂⟦ x ⟧ = [ff] ff∷ a[ x ] + FuExpr→FF idFE = [ff] + FuExpr→FF (x ⋆FE x₁) = (FuExpr→FF x) ff++ (FuExpr→FF x₁) + FuExpr→FF (invFE invF x) = ffInv invF (FuExpr→FF x) + FuExpr→FF ⟅ T , F , x ⟆FE = ff⟪ F , FuExpr→FF x ⟫ + + evAtom : {ws : wildStr} {x y : ob ws} + → FuAtom x y → ws H[ x , y ] + evAtom a[ x ] = x + evAtom {ws} a[ x , x₁ ]⁻ = inv ws x x₁ + + + invFuAtom-involR : ∀ ws {x y} invF → + (a : FuAtom {ws} y x) → + (ws ⟨ evAtom a ⋆ evAtom (invFuAtom invF a) ⟩) ≡ id ws + invFuAtom-involR ws invF a[ x ] = ⋆InvR ws _ _ + invFuAtom-involR ws invF a[ x , x₁ ]⁻ = ⋆InvL ws _ _ + + + evFF : {ws : wildStr} {x y : ob ws} + → FuFlat x y → ws H[ x , y ] + evFF {ws} [ff] = id ws + evFF {ws} (x ff∷ x₁) = ws ⟨ (evFF x) ⋆ (evAtom x₁) ⟩ + evFF (x invol invF ∷ff x₁) = evFF x + + + + evFF≡↓ : (ws : wildStr) {x y : ob ws} + → (f : FuFlat x y) → + evFF (ff↓ f) ≡ evFF f + evFF≡↓ ws [ff] = refl + evFF≡↓ ws (f ff∷ x) = cong (ws ⟨_⋆ _ ⟩) (evFF≡↓ ws f) + evFF≡↓ ws (f invol invF ∷ff x) = + ⋆Assoc ws _ _ _ + ∙∙ cong₂ (ws ⟨_⋆_⟩) (evFF≡↓ ws f) (invFuAtom-involR ws invF x) + ∙∙ ⋆IdR ws (evFF f) + + evFF++ : ∀ {ws x y z} → (g : FuFlat {ws} x y) → (h : FuFlat y z) → + ws ⟨ evFF g ⋆ evFF h ⟩ ≡ (evFF (g ff++ h)) + evFF++ {ws = ws} g [ff] = ⋆IdR ws _ + evFF++ {ws = ws} g (h ff∷ x) = sym (⋆Assoc ws _ _ _) ∙ + cong (ws ⟨_⋆ (evAtom x) ⟩) (evFF++ g h) + evFF++ g (h invol invF ∷ff x) = evFF++ g h + + evAinv : ∀ {ws x y} → (invF : InvFlag) → + (g : FuAtom {ws} x y) → + inv ws invF (evAtom g) ≡ evAtom (invFuAtom invF g) + evAinv invF a[ x ] = refl + evAinv {ws = ws} invF a[ x , x₁ ]⁻ with mbIsWildGroupoid | invol-inv ws + ... | just x₂ | ii = ii _ x₁ + + + aa-Func : ∀ {ws ws' x y} (F : WildFunctor (toWildCat ws) (toWildCat ws')) + (a : FuAtom {ws} x y) → + F-hom F (evAtom a) ≡ evAtom aa⟪ F , a ⟫ + aa-Func F a[ x ] = refl + aa-Func {ws} {ws'} F a[ invF , x₁ ]⁻ = + F-inv' (WG ws invF) (WG ws' invF) F x₁ + + aa-Func-inv : ∀ {ws ws' x y} invF (F : WildFunctor (toWildCat ws) (toWildCat ws')) + (a : FuAtom {ws} x y) → + inv ws' + invF (evAtom aa⟪ F , a ⟫) ≡ evAtom aa⟪ F , invFuAtom invF a ⟫ + aa-Func-inv invF F a[ x ] = refl + aa-Func-inv {ws' = ws'} invF F a[ x , x₁ ]⁻ with mbIsWildGroupoid | invol-inv ws' + ... | just x₂ | ii = ii _ _ + + evFFinv : ∀ {ws x y} → (invF : InvFlag) → + (g : FuFlat {ws} x y) → + inv ws invF (evFF g) ≡ evFF (ffInv invF g) + evFFinv {ws} invF [ff] = id≡inv-id ws invF + evFFinv {ws} invF (g ff∷ x) = + distInv ws invF _ _ + ∙∙ cong₂ (ws ⟨_⋆_⟩) (evAinv invF x ∙ sym (⋆IdL ws _)) (evFFinv invF g) + ∙∙ evFF++ _ (ffInv invF g) + evFFinv {ws} invF (g invol invF₁ ∷ff x) = + evFFinv invF g ∙∙ sym (⋆IdL ws _) ∙∙ evFF++ _ (ffInv invF g) + + ff⟪⟫⋆ : ∀ {ws ws' x y z} (F : WildFunctor (toWildCat ws) (toWildCat ws')) + → (f : FuFlat {ws} x y) → (g : FuFlat y z) → + ws' ⟨ (evFF ff⟪ F , f ⟫) ⋆ (evFF ff⟪ F , g ⟫) ⟩ + ≡ evFF ff⟪ F , f ff++ g ⟫ + ff⟪⟫⋆ {ws' = ws'} F f [ff] = ⋆IdR ws' _ + ff⟪⟫⋆ {ws' = ws'} F f (g ff∷ x) = + sym (⋆Assoc ws' _ _ _) ∙ + cong (ws' ⟨_⋆ (evAtom aa⟪ F , x ⟫)⟩) (ff⟪⟫⋆ F f g) + ff⟪⟫⋆ F f (g invol invF ∷ff x) = ff⟪⟫⋆ F f g + + ff⟪⟫inv : ∀ {ws ws' x y} invF (F : WildFunctor (toWildCat ws) (toWildCat ws')) + → (f : FuFlat {ws} x y) → + inv ws' invF (evFF ff⟪ F , f ⟫) + ≡ evFF ff⟪ F , ffInv invF f ⟫ + ff⟪⟫inv {ws' = ws'} invF F [ff] = id≡inv-id ws' invF + ff⟪⟫inv {ws' = ws'} invF F (f ff∷ x) = + distInv ws' invF _ _ + ∙∙ cong₂ (ws' ⟨_⋆_⟩) (aa-Func-inv invF F x ∙ sym (⋆IdL ws' _)) (ff⟪⟫inv invF F f) + ∙∙ ff⟪⟫⋆ F _ (ffInv invF f) + ff⟪⟫inv {ws' = ws'} invF F (f invol invF₁ ∷ff x) = + ff⟪⟫inv invF F f + ∙∙ sym (⋆IdL ws' _) + ∙∙ ff⟪⟫⋆ F ([ff] invol invF₁ ∷ff x) (ffInv invF f) + + + + ff-Func : ∀ {ws ws' x y} (F : WildFunctor (toWildCat ws) (toWildCat ws')) + (f : FuFlat {ws} x y) → + F-hom F (evFF f) ≡ evFF ff⟪ F , f ⟫ + ff-Func F [ff] = F-id F + ff-Func {ws' = ws'} F (f ff∷ x) = + F-seq F _ _ ∙ cong₂ (ws' ⟨_⋆_⟩) (ff-Func F f) (aa-Func F x) + ff-Func F (f invol invF ∷ff x) = ff-Func F f + + evFF-Func : ∀ {ws ws'} (F : WildFunctor (toWildCat ws) (toWildCat ws')) {x y} → + (f : FuExpr ws x y) → + F-hom F (evFuExpr f) ≡ evFF {ws'} ff⟪ F , FuExpr→FF f ⟫ + evFF-Func {ws' = ws'} F FuExpr'.𝒂⟦ x ⟧ = sym (⋆IdL ws' _) + evFF-Func F FuExpr'.idFE = F-id F + evFF-Func {ws} {ws'} F (f FuExpr'.⋆FE f₁) = + F-seq F _ _ ∙∙ cong₂ (ws' ⟨_⋆_⟩) (evFF-Func {ws} F f) ((evFF-Func {ws} F f₁)) + ∙∙ ff⟪⟫⋆ F (FuExpr→FF f) (FuExpr→FF f₁) + evFF-Func {ws} {ws'} F (FuExpr'.invFE invF f) = + F-inv' (WG ws invF) + (WG ws' invF) F (evFuExpr f) + ∙∙ cong (inv ws' invF) (evFF-Func F f) ∙∙ ff⟪⟫inv invF F (FuExpr→FF f) + evFF-Func F FuExpr'.⟅ T , F' , f ⟆FE = + cong (F-hom F) (evFF-Func F' f) ∙ + ff-Func F ff⟪ F' , FuExpr→FF f ⟫ + + evFuExpr≡evFF : (ws : wildStr) {x y : ob ws} + → (f : FuExpr ws x y) → + evFuExpr f ≡ evFF (FuExpr→FF f) + evFuExpr≡evFF ws FuExpr'.𝒂⟦ x ⟧ = sym (⋆IdL ws _) + evFuExpr≡evFF _ FuExpr'.idFE = refl + evFuExpr≡evFF ws (f FuExpr'.⋆FE f₁) = + cong₂ (ws ⟨_⋆_⟩) (evFuExpr≡evFF ws f) (evFuExpr≡evFF ws f₁) + ∙ evFF++ (FuExpr→FF f) (FuExpr→FF f₁) + evFuExpr≡evFF ws (FuExpr'.invFE invF f) = + cong (inv ws invF) (evFuExpr≡evFF ws f) ∙ evFFinv invF (FuExpr→FF f) + evFuExpr≡evFF ws FuExpr'.⟅ T , F , f ⟆FE = evFF-Func F f + + -- evFuExpr-singl : {ws : wildStr} {x y : ob ws} + -- → (f : FuExpr ws x y) → + -- Σ {!!} {!λ f' → evFuExpr f ≡ evFF (FuExpr→FF f)!} + -- evFuExpr-singl = {!!} + + + + magicNumber : ℕ + magicNumber = 100 + + infixl 5 us∷_ + + -- us∷_ : R.Term + us∷_ : List (R.Arg R.Term) → List (R.Arg R.Term) + us∷_ = R.unknown v∷_ + + buildFromTE : ∀ {W} → TE InvFlag (lift W) _ _ → R.Term + buildFromTE FuExpr'.𝒂⟦ x ⟧ = R.con (quote FuExpr'.𝒂⟦_⟧) ([ varg x ]) + buildFromTE FuExpr'.idFE = R.con (quote FuExpr'.idFE) [] + buildFromTE (x FuExpr'.⋆FE x₁) = + R.con (quote FuExpr'._⋆FE_) + (buildFromTE x v∷ buildFromTE x₁ v∷ []) + buildFromTE (FuExpr'.invFE invF x) = + R.con (quote FuExpr'.invFE) + (us∷ buildFromTE x v∷ []) + buildFromTE FuExpr'.⟅ lift T , F , x ⟆FE = + R.con (quote FuExpr'.⟅_,_,_⟆FE) + (T v∷ F v∷ buildFromTE x v∷ []) + + module tryWCE WS (tGs : List R.Term) where + + + mb-invol : R.Term → ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) + mb-invol _ zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) + mb-invol _ _ (R.con (quote [ff]) _) = R.returnTC nothing + mb-invol W (suc n) (R.con (quote _ff∷_) tl) = match2Vargs tl >>= w + where + w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) + w (R.con (quote [ff]) _ , _) = R.returnTC nothing + w (xs'@(R.con (quote _ff∷_) tl') , y) = + match2Vargs tl' >>= λ (xs , x) → + R.catchTC + (R.noConstraints $ R.unify + (R.def (quote invFuAtomExplicit) (WS v∷ W v∷ us∷ x v∷ [])) y + >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol W n xs) + >>= λ (xs* , xs*↑) → + R.returnTC + (just ((R.con (quote _invol_∷ff_) (xs* v∷ us∷ x v∷ [])) + , xs*) + )) + (map-Maybe (map-both (λ xs* → R.con (quote _ff∷_) + ((xs* v∷ y v∷ [])))) + <$> mb-invol W n xs') + w (x , y) = R.typeError ("Imposible!! : " ∷ₑ x ∷ₑ " " ∷ₑ y ∷ₑ []) + mb-invol _ _ x = R.typeError ("Imposible!! : " ∷ₑ x ∷ₑ []) + + mb-invol' : R.Term → R.Term → R.TC (Maybe (R.Term × R.Term)) + mb-invol' = λ W → mb-invol W magicNumber + + + redList : R.Term → ℕ → R.Term → R.TC (List R.Term) + redList W = h + where + h : ℕ → R.Term → R.TC (List R.Term) + h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) + h (suc k) t = + (mb-invol' W t) >>= + Mb.rec + (R.returnTC []) + λ (t' , t'↓) → do + p' ← h k t'↓ + -- R.typeError [ R.termErr t' ] + R.returnTC (t' ∷ p') + + + redList' : R.Term → R.Term → R.TC (List R.Term) + redList' W = redList W magicNumber + + -- involPathTerm : R.Term → ℕ → R.Term → R.TC R.Term + -- involPathTerm W (suc n) t = + -- (mb-invol' W t) >>= + -- Mb.rec + -- (R.returnTC (R.def (quote refl) [])) + -- λ (t' , t'↓) → do + -- p' ← involPathTerm W n t'↓ + -- R.returnTC $ R.def (quote _∙_) + -- ((R.def (quote evFF≡↓) (WS v∷ W v∷ t' v∷ [])) v∷ p' v∷ []) + -- involPathTerm _ zero _ = + -- R.typeError ("magic number too low in mb-invol" ∷ₑ []) + + -- involPathTerm' : R.Term → R.Term → R.TC R.Term + -- involPathTerm' W t = involPathTerm W magicNumber t + -- -- R.checkType pa (R.def (quote Path) {!!}) + + checkFromTE : ∀ {W} → TE InvFlag (lift W) _ _ → + R.TC R.Term --(Σ wildStr λ ws → Σ _ (λ x → Σ _ (FuExpr ws x))) + checkFromTE {W} te = do + let te' = buildFromTE te + R.checkType te' + (R.def (quote FuExpr) (WS v∷ W v∷ us∷ us∷ [] )) + -- (R.con (quote _,_) + -- (W + -- v∷ R.con (quote _,_) + -- (R.unknown v∷ R.con (quote _,_) + -- (R.unknown v∷ te' v∷ []) v∷ []) v∷ []) ) + -- ? + + -- if = {!!} + + tryE : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) + + fromWC : R.Term → R.TC R.Term + fromWC t = tryAllTC + (R.typeError ("fromWC fail: " ∷ₑ t ∷ₑ [])) + tGs + λ ws → R.unify (R.def (quote toWildCat) + (WS v∷ ws v∷ [])) t >> R.returnTC ws + + tryOp : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) + tryOp W zero _ = R.typeError [] + tryOp W (suc k) t = do + let tm = R.con (quote FuCases._⋆FE_) + (R.unknown v∷ R.unknown v∷ []) + ty = R.def (quote FuCases) + (WS v∷ W v∷ t v∷ []) + tm' ← R.checkType tm ty + (t1 , t2) ← h tm' + t1' ← tryE W k t1 + t2' ← tryE W k t2 + R.returnTC (t1' TermExpr.⋆FE t2') + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con _ l) = match2Vargs l + h t = R.typeError [] + + tryInv : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) + tryInv W zero _ = R.typeError [] + tryInv W (suc k) t = do + let tm = R.con (quote FuCases.invFE) + (R.unknown v∷ R.unknown v∷ []) + ty = R.def (quote FuCases) + (WS v∷ W v∷ t v∷ []) + tm' ← R.checkType tm ty + (_ , t-x) ← h tm' + t-x' ← tryE W k t-x + ifQ ← R.unquoteTC R.unknown + R.returnTC (TermExpr.invFE ifQ t-x') + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con _ l) = match2Vargs l + h t = R.typeError [] + + + tryFunc : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) + tryFunc W zero _ = R.typeError [] + tryFunc W (suc k) t = do + t' ← R.normalise t + -- (R.typeError $ "tryFunc fail " ∷nl t ∷nl t' ∷nl getConTail t') + (WC-src , F-t , x-t) ← Mb.rec + (R.typeError $ "tryFunc fail " ∷nl t ∷nl t' ∷nl getConTail t') + (λ (F-t , x-t) → do + F-ty ← R.withNormalisation true $ R.inferType F-t + (W-src , W-trg) ← h F-ty + R.returnTC {A = R.Term × R.Term × R.Term} + (W-src , (F-t , x-t)) + ) + (mbFunctorApp t') + WS-src ← fromWC WC-src + let tm = R.con (quote FuCases.⟅_,_,_⟆FE) + (WS-src v∷ F-t v∷ x-t v∷ []) + ty = R.def (quote FuCases) + (WS v∷ W v∷ t v∷ []) + tm' ← R.checkType tm ty + x-t' ← tryE WS-src k x-t + R.returnTC (TermExpr.⟅ lift WS-src , F-t , x-t' ⟆FE) + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con _ l) = match2Vargs l + h (R.def _ l) = match2Vargs l + h t = R.typeError $ "match2Fail: " ∷ₑ t ∷ₑ [] + + + + tryId : (W : R.Term) → R.Term → R.TC (TE InvFlag (lift W) _ _) + tryId W t = do + let tm = R.con (quote FuCases.idFE) [] + ty = R.def (quote FuCases) + (WS v∷ W v∷ t v∷ []) + tm' ← R.checkType tm ty + R.returnTC (TermExpr.idFE) + + atom : (W : R.Term) → R.Term → R.TC (TE InvFlag (lift W) _ _) + atom _ x = R.returnTC $ TermExpr.𝒂⟦ x ⟧ + + + tryE W zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryE W (suc k) t = + R.catchTC + (tryId W t) + (R.catchTC (tryInv W k t) + -- (tryOp k t) + (R.catchTC (tryOp W k t) + (R.catchTC (tryFunc W k t) (atom W t))) + ) + + + + solveW : R.Term → R.Term → R.Term → R.TC Unit + solveW Ws Wts' hole = do + Wts ← quotedList→ListOfTerms Wts' + Wt ← tryAllTC + (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) + Wts R.returnTC + hoTy ← R.withNormalisation true $ + R.inferType hole >>= wait-for-type + (t0 , t1) ← (get-boundary hoTy ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 + t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 + expr0 ← tryWCE.checkFromTE Ws Wts t0' + expr1 ← tryWCE.checkFromTE Ws Wts t1' + -- R.typeError + let msg = (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl + TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) + invol0 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr0 ])) + invol1 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr1 ])) + + + red0 ← tryWCE.redList' Ws Wts Wt invol0 + red1 ← tryWCE.redList' Ws Wts Wt invol1 + + let invPa0 = Li.map + (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) + red0 + invPa1 = Li.map + (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) + red1 + -- R.typeError ( + -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red0)) + -- ++ₑ " " ∷nl " " ∷nl + -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red1)) ++ₑ []) + +-- -- invPa0 ← tryWCE.involPathTerm' Ws Wts Wt invol0 -->>= R.normalise +-- -- invPa1 ← tryWCE.involPathTerm' Ws Wts Wt invol1 -->>= R.normalise + +-- -- -- R.typeError (invol0 ∷nl +-- -- -- invol1 ∷ₑ []) + + + + let finalTrm0 = + just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) + ∷ invPa0 + -- R.def (quote _∙_) + -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) v∷ + -- invPa0 v∷ []) + finalTrm1 = + just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) + ∷ invPa1 + -- R.def (quote _∙_) + -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) v∷ + -- invPa1 v∷ []) + + let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms + (finalTrm0 ++ symPathTerms finalTrm1) + -- --invPa0 + -- R.def (quote _∙_) + -- (finalTrm0 v∷ --R.unknown v∷ + -- (R.def (quote sym) + -- (finalTrm1 v∷ [])) v∷ []) + + -- msg ← R.formatErrorParts (finalTrm ∷ₑ []) + + -- R.typeError (msg ∷ₑ []) + + (R.unify finalTrm hole) -- <|> R.typeError msg + + +-- macro + + +-- testInferExpr : R.Term → R.Term → R.Term → R.Term → R.TC Unit +-- testInferExpr Ws Wts' t ho = do +-- Wts ← quotedList→ListOfTerms Wts' +-- Wt ← tryAllTC +-- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) +-- Wts R.returnTC +-- expr ← tryWCE.tryE Ws Wts Wt magicNumber t -->> R.returnTC _ + +-- tryWCE.checkFromTE Ws Wts expr +-- R.typeError (TermExpr.printFuExpr InvFlag (λ _ → "●") expr ∷ₑ []) +-- -- R.unify R.unknown ho + + +-- -- solveNoInvs : R.Term → R.Term → R.Term → R.TC Unit +-- -- solveNoInvs Ws Wts' hole = do +-- -- Wts ← quotedList→ListOfTerms Wts' +-- -- Wt ← tryAllTC +-- -- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) +-- -- Wts R.returnTC +-- -- (t0 , t1) ← R.withNormalisation true $ +-- -- R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec +-- -- (R.typeError [ R.strErr "unable to get boundary" ]) +-- -- (λ x → R.returnTC x) +-- -- t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 +-- -- t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 +-- -- expr0 ← tryWCE.checkFromTE Ws Wts t0' +-- -- expr1 ← tryWCE.checkFromTE Ws Wts t1' +-- -- -- R.typeError +-- -- -- R.typeError (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl +-- -- -- TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) + +-- -- let finalTrm = R.def (quote _∙∙_∙∙_) +-- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ expr0 v∷ [])) v∷ +-- -- (R.def (quote refl) []) v∷ +-- -- (R.def (quote sym) +-- -- ((R.def (quote evFuExpr≡evFF)) (Ws v∷ expr1 v∷ []) v∷ [])) v∷ []) +-- -- R.unify finalTrm hole + + +-- solveW : R.Term → R.Term → R.Term → R.TC Unit +-- solveW Ws Wts' hole = do +-- Wts ← quotedList→ListOfTerms Wts' +-- Wt ← tryAllTC +-- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) +-- Wts R.returnTC +-- hoTy ← R.withNormalisation true $ +-- R.inferType hole >>= wait-for-type +-- (t0 , t1) ← (get-boundary hoTy ) >>= Mb.rec +-- (R.typeError [ R.strErr "unable to get boundary" ]) +-- (λ x → R.returnTC x) +-- t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 +-- t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 +-- expr0 ← tryWCE.checkFromTE Ws Wts t0' +-- expr1 ← tryWCE.checkFromTE Ws Wts t1' +-- -- R.typeError +-- let msg = (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl +-- TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) +-- invol0 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr0 ])) +-- invol1 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr1 ])) + + +-- red0 ← tryWCE.redList' Ws Wts Wt invol0 +-- red1 ← tryWCE.redList' Ws Wts Wt invol1 + +-- let invPa0 = Li.map +-- (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) +-- red0 +-- invPa1 = Li.map +-- (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) +-- red1 +-- -- R.typeError ( +-- -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red0)) +-- -- ++ₑ " " ∷nl " " ∷nl +-- -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red1)) ++ₑ []) + +-- -- -- invPa0 ← tryWCE.involPathTerm' Ws Wts Wt invol0 -->>= R.normalise +-- -- -- invPa1 ← tryWCE.involPathTerm' Ws Wts Wt invol1 -->>= R.normalise + +-- -- -- -- R.typeError (invol0 ∷nl +-- -- -- -- invol1 ∷ₑ []) + + + +-- let finalTrm0 = +-- just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) +-- ∷ invPa0 +-- -- R.def (quote _∙_) +-- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) v∷ +-- -- invPa0 v∷ []) +-- finalTrm1 = +-- just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) +-- ∷ invPa1 +-- -- R.def (quote _∙_) +-- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) v∷ +-- -- invPa1 v∷ []) + +-- let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms +-- (finalTrm0 ++ symPathTerms finalTrm1) +-- -- --invPa0 +-- -- R.def (quote _∙_) +-- -- (finalTrm0 v∷ --R.unknown v∷ +-- -- (R.def (quote sym) +-- -- (finalTrm1 v∷ [])) v∷ []) + +-- -- msg ← R.formatErrorParts (finalTrm ∷ₑ []) + +-- -- R.typeError (msg ∷ₑ []) + +-- (R.unify finalTrm hole) -- <|> R.typeError msg + + + +-- module _ (G : Group ℓ) where +-- open GroupStr (snd G) +-- Group→WildGroupoid : WildGroupoid ℓ-zero ℓ +-- WildCat.ob (WildGroupoid.wildCat Group→WildGroupoid) = Unit +-- WildCat.Hom[_,_] (WildGroupoid.wildCat Group→WildGroupoid) _ _ = ⟨ G ⟩ +-- WildCat.id (WildGroupoid.wildCat Group→WildGroupoid) = 1g +-- WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ +-- WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL +-- WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR +-- WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) +-- wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f +-- wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f +-- wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f + + +-- GroupWS : WildStr ℓ-zero ℓ +-- GroupoidWS WildCatWS : WildStr ℓ ℓ' + +-- WildStr.wildStr (GroupWS {ℓ}) = Group ℓ +-- WildStr.toWildCat (GroupWS) = WildGroupoid.wildCat ∘ Group→WildGroupoid +-- WildStr.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) + +-- WildStr.wildStr (GroupoidWS {ℓ} {ℓ'}) = WildGroupoid ℓ ℓ' +-- WildStr.toWildCat GroupoidWS = WildGroupoid.wildCat +-- WildStr.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid + + + diff --git a/Cubical/WildCat/Functor.agda b/Cubical/WildCat/Functor.agda index 05b363b02c..ba579d5d36 100644 --- a/Cubical/WildCat/Functor.agda +++ b/Cubical/WildCat/Functor.agda @@ -167,9 +167,37 @@ WildFunctor.F-seq commFunctor _ _ = refl module _ {C : WildGroupoid ℓC ℓC'} {D : WildGroupoid ℓD ℓD'} (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where - - module gC = WildGroupoid C - module gD = WildGroupoid D + + private + module gC = WildGroupoid C + module gD = WildGroupoid D F-inv : ∀ {x y} f → F-hom F (gC.inv {x} {y} f) ≡ gD.inv (F-hom F f) F-inv f = gD.invUniqueR $ sym (F-seq F _ _) ∙∙ congS (F-hom F) (gC.⋆InvR f) ∙∙ F-id F + +module _ + (C : WildGroupoid ℓC ℓC') (D : WildGroupoid ℓD ℓD') + (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where + + private + module gC = WildGroupoid C + module gD = WildGroupoid D + + F-inv' : ∀ {x y} f → F-hom F (gC.inv {x} {y} f) ≡ gD.inv (F-hom F f) + F-inv' f = gD.invUniqueR $ sym (F-seq F _ _) ∙∙ congS (F-hom F) (gC.⋆InvR f) ∙∙ F-id F + + +-- action on objects +infix 30 _⟅_⟆ +_⟅_⟆ : {C : WildCat ℓC ℓC'} {D : WildCat ℓD ℓD'} (F : WildFunctor C D) + → C .ob + → D .ob +_⟅_⟆ = F-ob + +-- action on morphisms +infix 30 _⟪_⟫ -- same infix level as on objects since these will never be used in the same context +_⟪_⟫ : {C : WildCat ℓC ℓC'} {D : WildCat ℓD ℓD'} (F : WildFunctor C D) + → ∀ {x y} + → C [ x , y ] + → D [(F ⟅ x ⟆) , (F ⟅ y ⟆)] +_⟪_⟫ = F-hom From 70c6ed7f9e780effbbb6d79ed202c470fa4872d2 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 16:30:03 +0100 Subject: [PATCH 04/20] working, wip --- Cubical/Tactics/GroupSolver/Example.agda | 64 ++++++++ Cubical/Tactics/GroupSolver/Solver.agda | 101 ++++++++++++ Cubical/Tactics/GroupoidSolver/Example.agda | 172 +++++--------------- Cubical/Tactics/GroupoidSolver/Solver.agda | 19 +-- Cubical/Tactics/WildCatSolver/Example.agda | 129 --------------- Cubical/Tactics/WildCatSolver/Solver.agda | 1 - Cubical/Tactics/WildCatSolver/Solvers.agda | 2 +- 7 files changed, 217 insertions(+), 271 deletions(-) create mode 100644 Cubical/Tactics/GroupSolver/Example.agda create mode 100644 Cubical/Tactics/GroupSolver/Solver.agda diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda new file mode 100644 index 0000000000..44c4aa3007 --- /dev/null +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupSolver.Example where + +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.Group +open import Cubical.Tactics.GroupSolver.Solver +open import Cubical.Data.List + +open import Cubical.WildCat.Functor + +private + variable + ℓ : Level + + + +module example (G G* G○ : Group ℓ) + (F* : GroupHom' G* G) + (F○ : GroupHom' G○ G*) + where + + + open Group-Solver ℓ + + + open GroupStr (snd G) + + open WildFunctor + + module * where + open GroupStr (snd G*) public + + module ○ where + open GroupStr (snd G○) public + + + + module T1 (p p' q r : fst G ) where + + + pA pB : fst G + pA = ((((((1g · p') · q) · (inv q)) · (inv p')) · p) · q) · r + pB = ((1g · p) · q) · r + + pA≡pB : pA ≡ pB + pA≡pB = solveGroup (G ∷ []) + + + + module T2 p p' q r s t u where + + + lhs rhs : fst G + lhs = (p · p') · (inv p' · (F* ⟪ (((*.inv q) *.· r) *.· F○ ⟪ ○.inv t ⟫ *.· + (*.inv (F○ ⟪ s ○.· s ⟫) *.· F○ ⟪ u ⟫ )) ⟫ )) + + rhs = inv (F* ⟪ q ⟫ · inv p) · (F* ⟪ r ⟫) · + (comp-WildFunctor F○ F*) ⟪ ○.inv (s ○.· s ○.· t) ○.· u ⟫ + + + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveGroup (G ∷ G* ∷ G○ ∷ []) diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda new file mode 100644 index 0000000000..2233ef47b7 --- /dev/null +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -0,0 +1,101 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.GroupSolver.Solver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function as Fu +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +-- open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + +-- open import Cubical.WildCat.WGE +open import Cubical.WildCat.Base +open import Cubical.WildCat.Functor +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties + +open import Cubical.Tactics.WildCatSolver.Solvers + + +module _ {ℓ} where + + module _ (G : Group ℓ) where + open GroupStr (snd G) + Group→WildGroupoid : WildGroupoid ℓ-zero ℓ + WildCat.ob (WildGroupoid.wildCat Group→WildGroupoid) = Unit + WildCat.Hom[_,_] (WildGroupoid.wildCat Group→WildGroupoid) _ _ = ⟨ G ⟩ + WildCat.id (WildGroupoid.wildCat Group→WildGroupoid) = 1g + WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ + WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL + WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR + WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) + wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f + wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f + wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f + + + GroupHom' : (G H : Group ℓ) → Type ℓ + + GroupHom' G H = WildFunctor + (WildGroupoid.wildCat (Group→WildGroupoid G)) + (WildGroupoid.wildCat (Group→WildGroupoid H)) + + IsoGroupHom' : ∀ {G H} → Iso (GroupHom' G H) (GroupHom G H) + Iso.fun IsoGroupHom' wf = _ , makeIsGroupHom (WildFunctor.F-seq wf) + WildFunctor.F-ob (Iso.inv IsoGroupHom' _) = λ _ → tt + WildFunctor.F-hom (Iso.inv IsoGroupHom' (f , _)) = f + WildFunctor.F-id (Iso.inv IsoGroupHom' (_ , gh)) = IsGroupHom.pres1 gh + WildFunctor.F-seq (Iso.inv IsoGroupHom' (_ , gh)) = IsGroupHom.pres· gh + Iso.rightInv IsoGroupHom' _ = GroupHom≡ refl + WildFunctor.F-ob (Iso.leftInv IsoGroupHom' _ i) = λ _ → tt + WildFunctor.F-hom (Iso.leftInv IsoGroupHom' wf i) = WildFunctor.F-hom wf + WildFunctor.F-id (Iso.leftInv (IsoGroupHom' {G = G} {H = H}) wf i) = + IsGroup.is-set (GroupStr.isGroup (snd H)) + (WildFunctor.F-hom wf (GroupStr.1g (snd G))) + (GroupStr.1g (snd H)) + (hom1g (G .snd) (WildFunctor.F-hom wf) (H .snd) + (WildFunctor.F-seq wf)) + (WildFunctor.F-id wf) i + WildFunctor.F-seq (Iso.leftInv IsoGroupHom' wf i) = λ f g → WildFunctor.F-seq wf f g + + +module Group-Solver ℓ where + + GroupWS : WildStr ℓ-zero ℓ + WildStr.wildStr GroupWS = Group ℓ + WildStr.toWildCat GroupWS = WildGroupoid.wildCat ∘ Group→WildGroupoid + WildStr.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) + + private + module GRP-WS = WildStr GroupWS + + macro + solveGroup : R.Term → R.Term → R.TC Unit + solveGroup = GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) + diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 8000876cfa..583c4cfe02 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -1,12 +1,13 @@ {-# OPTIONS --safe #-} -module Cubical.Tactics.WildCatSolver.Example where +module Cubical.Tactics.GroupoidSolver.Example where open import Cubical.Foundations.Prelude open import Cubical.WildCat.Base -open import Cubical.Tactics.WildCatSolver.Solver +open import Cubical.Tactics.GroupoidSolver.Solver open import Cubical.Data.List +open import Cubical.Data.Nat open import Cubical.WildCat.Functor @@ -14,157 +15,68 @@ private variable ℓ ℓ' : Level -module exampleC (WC WC* : WildCat ℓ ℓ') - (F : WildFunctor WC* WC) where - open WildCat WC - module * = WildCat WC* - - module _ x y z w v - (p : Hom[ x , F ⟅ y ⟆ ]) - (q : *.Hom[ y , z ]) - (r : *.Hom[ z , w ]) - (s : Hom[ F ⟅ w ⟆ , v ]) where - - - pA pB pC : Hom[ x , v ] - pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) - pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) - pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) - - open WildCat-Solver ℓ ℓ' - - pA=pB : pA ≡ pB - pA=pB = solveWildCat (WC ∷ WC* ∷ []) - - pB=pC : pB ≡ pC - pB=pC = solveWildCat (WC ∷ WC* ∷ []) - - - --- -- module example (WG WG* : WildGroupoid ℓ ℓ') --- -- (F : WildFunctor --- -- (WildGroupoid.wildCat WG*) --- -- (WildGroupoid.wildCat WG)) --- -- where - --- -- module WildStrGPD = WildStr (GroupoidWS {ℓ} {ℓ'}) - --- -- -- open WildStrGPD using --- -- -- (testInferExpr; solveNoInvs; solveW) - --- -- open WildGroupoid WG --- -- open WildCat wildCat - --- -- open WildFunctor - - --- -- module * where --- -- open WildCat (WildGroupoid.wildCat WG*) public --- -- open WildGroupoid WG* public - --- -- -- module T1 x y z w --- -- -- (p : Hom[ x , y ]) --- -- -- (q : Hom[ y , z ]) --- -- -- (r : Hom[ z , w ]) where +module example (WG WG* : WildGroupoid ℓ ℓ') + (F : WildFunctor + (WildGroupoid.wildCat WG*) + (WildGroupoid.wildCat WG)) + where --- -- -- pA pB : Hom[ x , w ] --- -- -- pA = p ⋆ (q ⋆ r) --- -- -- pB = (p ⋆ q) ⋆ r --- -- -- -- pF : + open WildGroupoid-Solver ℓ ℓ' --- -- -- pXX : Hom[ x , x ] --- -- -- pXX = id + open WildGroupoid WG + open WildCat wildCat --- -- -- tryTestOp : Unit --- -- -- tryTestOp = testTryOp GroupoidWS WG pA + open WildFunctor --- -- -- tryTestId : Unit --- -- -- tryTestId = testTryId GroupoidWS WG pXX --- -- -- pACase : FuCases WG pA --- -- -- pACase = p FuCases.⋆FE (q ⋆ r) + module * where + open WildCat (WildGroupoid.wildCat WG*) public + open WildGroupoid WG* public --- -- module T1 x y z w --- -- (p p' : Hom[ x , y ]) --- -- (q : Hom[ y , z ]) --- -- (r : Hom[ z , w ]) where + module T1 x y z w + (p p' : Hom[ x , y ]) + (q : Hom[ y , z ]) + (r : Hom[ z , w ]) where --- -- -- pA pB : Hom[ x , x ] --- -- -- pA = (p ⋆ inv p) --- -- -- pB = id + pA pB : Hom[ x , w ] + pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r + pB = ((id ⋆ p) ⋆ q) ⋆ r --- -- -- pA≡pB : pA ≡ pB --- -- -- pA≡pB = {!solveW GroupoidWS (WG ∷ WG* ∷ [])!} - - --- -- pA pB : Hom[ x , w ] --- -- pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r --- -- -- pA = ((((((id ⋆ p))) ⋆ (inv p)) ⋆ p) ⋆ q) ⋆ r --- -- pB = ((id ⋆ p) ⋆ q) ⋆ r - --- -- pA≡pB : pA ≡ pB --- -- pA≡pB = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) + pA≡pB : pA ≡ pB + pA≡pB = solveWildGroupoid (WG ∷ WG* ∷ []) --- -- module T2 x y z w --- -- (p : Hom[ x , F ⟅ y ⟆ ]) --- -- (p' : Hom[ F ⟅ y ⟆ , x ]) --- -- (q : *.Hom[ z , y ]) --- -- (r : *.Hom[ z , w ]) where - - --- -- pA : Hom[ F ⟅ y ⟆ , F-ob F w ] --- -- pA = F-hom F ((*.inv q) *.⋆ r) --- -- pB pB' pB'' : Hom[ x , F ⟅ w ⟆ ] --- -- pB = (p ⋆ (F-hom F ((*.inv q) *.⋆ r))) --- -- pB'' = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) --- -- pB' = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) - - --- -- pB''≡pB' : pB'' ≡ pB' --- -- pB''≡pB' = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) - - - - - - --- -- -- module example (WG WG* : WildGroupoid ℓ ℓ') --- -- -- (F : WildFunctor (WildGroupoid.wildCat WG*) (WildGroupoid.wildCat WG)) where --- -- -- open WildGroupoid WG --- -- -- open WildCat wildCat - - + module T2 x y z w + (p : Hom[ x , F ⟅ y ⟆ ]) + (p' : Hom[ F ⟅ y ⟆ , x ]) + (q : *.Hom[ z , y ]) + (r : *.Hom[ z , w ]) where --- -- -- module * where --- -- -- open WildCat (WildGroupoid.wildCat WG*) public --- -- -- open WildGroupoid WG* public --- -- -- module _ x v y z w --- -- -- (p p' : Hom[ x , F ⟅ y ⟆ ]) --- -- -- (q q' : *.Hom[ y , z ]) --- -- -- (r : *.Hom[ w , z ]) --- -- -- (s : Hom[ F ⟅ w ⟆ , v ]) where + lhs rhs : Hom[ x , F-ob F w ] + lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) + rhs = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) --- -- -- pA pB pC : Hom[ x , v ] --- -- -- pA = (p ⋆ (id ⋆ F ⟪ q *.⋆ *.inv q' ⟫)) ⋆ (F ⟪ q' *.⋆ (*.inv r) ⟫ ⋆ s) --- -- -- pB = ((p ⋆ F ⟪ q *.⋆ (*.inv (*.inv r *.⋆ r) *.⋆ *.inv q)⟫) --- -- -- ⋆ F ⟪ q *.⋆ *.id ⟫) ⋆ ((inv (F ⟪ r ⟫ ⋆ id)) ⋆ s) + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (WG ∷ WG* ∷ []) --- -- -- pC = p' ⋆ (inv p' ⋆ (p ⋆ (F ⟪ q *.⋆ (*.inv r) ⟫ ⋆ s))) --- -- -- pA=pB : pA ≡ pB --- -- -- pA=pB = solveGroupoid WG WG* + module T3 (obs : ℕ → ob) + (homs : ∀ x y → ℕ → Hom[ obs x , obs y ]) + where --- -- -- pA=pC : pA ≡ pC --- -- -- pA=pC = solveGroupoid WG WG* + lhs rhs : Hom[ obs 1 , obs 5 ] + lhs = homs _ 2 1 ⋆ (homs _ _ 2 ⋆ (homs 3 5 2 ⋆ (homs _ 6 3 ⋆ inv (homs _ 6 3)))) + rhs = ((homs _ 2 1 ⋆ homs _ _ 2) ⋆ id) ⋆ homs 3 5 2 + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (WG ∷ []) diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index f1c8259ece..727eff5981 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Tactics.Groupoid.Solver where +module Cubical.Tactics.GroupoidSolver.Solver where open import Cubical.Foundations.Prelude @@ -29,7 +29,6 @@ open import Cubical.Relation.Binary open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.GroupoidSolver.Reflection open import Cubical.Tactics.Reflection open import Agda.Builtin.String @@ -41,17 +40,17 @@ open import Cubical.Algebra.Group open import Cubical.Tactics.WildCatSolver.Solvers -module WildCat-Solver ℓ ℓ' where +module WildGroupoid-Solver ℓ ℓ' where - WildCatWS : WildStr ℓ ℓ' - WildStr.wildStr (WildCatWS) = WildCat ℓ ℓ' - WildStr.toWildCat WildCatWS = idfun _ - WildStr.mbIsWildGroupoid WildCatWS = nothing + GroupoidWS : WildStr ℓ ℓ' + WildStr.wildStr GroupoidWS = WildGroupoid ℓ ℓ' + WildStr.toWildCat GroupoidWS = WildGroupoid.wildCat + WildStr.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid private - module WC-WS = WildStr WildCatWS + module GPD-WS = WildStr GroupoidWS macro - solveWildCat : R.Term → R.Term → R.TC Unit - solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) + solveWildGroupoid : R.Term → R.Term → R.TC Unit + solveWildGroupoid = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index 8000876cfa..e982a978a4 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -39,132 +39,3 @@ module exampleC (WC WC* : WildCat ℓ ℓ') pB=pC : pB ≡ pC pB=pC = solveWildCat (WC ∷ WC* ∷ []) - - - --- -- module example (WG WG* : WildGroupoid ℓ ℓ') --- -- (F : WildFunctor --- -- (WildGroupoid.wildCat WG*) --- -- (WildGroupoid.wildCat WG)) --- -- where - --- -- module WildStrGPD = WildStr (GroupoidWS {ℓ} {ℓ'}) - --- -- -- open WildStrGPD using --- -- -- (testInferExpr; solveNoInvs; solveW) - --- -- open WildGroupoid WG --- -- open WildCat wildCat - --- -- open WildFunctor - - --- -- module * where --- -- open WildCat (WildGroupoid.wildCat WG*) public --- -- open WildGroupoid WG* public - --- -- -- module T1 x y z w --- -- -- (p : Hom[ x , y ]) --- -- -- (q : Hom[ y , z ]) --- -- -- (r : Hom[ z , w ]) where - - --- -- -- pA pB : Hom[ x , w ] --- -- -- pA = p ⋆ (q ⋆ r) --- -- -- pB = (p ⋆ q) ⋆ r - --- -- -- -- pF : - --- -- -- pXX : Hom[ x , x ] --- -- -- pXX = id - - --- -- -- tryTestOp : Unit --- -- -- tryTestOp = testTryOp GroupoidWS WG pA - --- -- -- tryTestId : Unit --- -- -- tryTestId = testTryId GroupoidWS WG pXX - --- -- -- pACase : FuCases WG pA --- -- -- pACase = p FuCases.⋆FE (q ⋆ r) - - --- -- module T1 x y z w --- -- (p p' : Hom[ x , y ]) --- -- (q : Hom[ y , z ]) --- -- (r : Hom[ z , w ]) where - - --- -- -- pA pB : Hom[ x , x ] --- -- -- pA = (p ⋆ inv p) --- -- -- pB = id - --- -- -- pA≡pB : pA ≡ pB --- -- -- pA≡pB = {!solveW GroupoidWS (WG ∷ WG* ∷ [])!} - - --- -- pA pB : Hom[ x , w ] --- -- pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r --- -- -- pA = ((((((id ⋆ p))) ⋆ (inv p)) ⋆ p) ⋆ q) ⋆ r --- -- pB = ((id ⋆ p) ⋆ q) ⋆ r - --- -- pA≡pB : pA ≡ pB --- -- pA≡pB = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) - - - --- -- module T2 x y z w --- -- (p : Hom[ x , F ⟅ y ⟆ ]) --- -- (p' : Hom[ F ⟅ y ⟆ , x ]) --- -- (q : *.Hom[ z , y ]) --- -- (r : *.Hom[ z , w ]) where - - --- -- pA : Hom[ F ⟅ y ⟆ , F-ob F w ] --- -- pA = F-hom F ((*.inv q) *.⋆ r) --- -- pB pB' pB'' : Hom[ x , F ⟅ w ⟆ ] --- -- pB = (p ⋆ (F-hom F ((*.inv q) *.⋆ r))) --- -- pB'' = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) --- -- pB' = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) - - --- -- pB''≡pB' : pB'' ≡ pB' --- -- pB''≡pB' = WildStrGPD.solveW GroupoidWS (WG ∷ WG* ∷ []) - - - - - - --- -- -- module example (WG WG* : WildGroupoid ℓ ℓ') --- -- -- (F : WildFunctor (WildGroupoid.wildCat WG*) (WildGroupoid.wildCat WG)) where --- -- -- open WildGroupoid WG --- -- -- open WildCat wildCat - - - --- -- -- module * where --- -- -- open WildCat (WildGroupoid.wildCat WG*) public --- -- -- open WildGroupoid WG* public - --- -- -- module _ x v y z w --- -- -- (p p' : Hom[ x , F ⟅ y ⟆ ]) --- -- -- (q q' : *.Hom[ y , z ]) --- -- -- (r : *.Hom[ w , z ]) --- -- -- (s : Hom[ F ⟅ w ⟆ , v ]) where - - --- -- -- pA pB pC : Hom[ x , v ] --- -- -- pA = (p ⋆ (id ⋆ F ⟪ q *.⋆ *.inv q' ⟫)) ⋆ (F ⟪ q' *.⋆ (*.inv r) ⟫ ⋆ s) --- -- -- pB = ((p ⋆ F ⟪ q *.⋆ (*.inv (*.inv r *.⋆ r) *.⋆ *.inv q)⟫) --- -- -- ⋆ F ⟪ q *.⋆ *.id ⟫) ⋆ ((inv (F ⟪ r ⟫ ⋆ id)) ⋆ s) - --- -- -- pC = p' ⋆ (inv p' ⋆ (p ⋆ (F ⟪ q *.⋆ (*.inv r) ⟫ ⋆ s))) - --- -- -- pA=pB : pA ≡ pB --- -- -- pA=pB = solveGroupoid WG WG* - --- -- -- pA=pC : pA ≡ pC --- -- -- pA=pC = solveGroupoid WG WG* - - diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index d92216c4c0..85bc64d0e3 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -29,7 +29,6 @@ open import Cubical.Relation.Binary open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.GroupoidSolver.Reflection open import Cubical.Tactics.Reflection open import Agda.Builtin.String diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index 7af64e3e02..e8f7b5f589 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -29,7 +29,7 @@ open import Cubical.Relation.Binary open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.GroupoidSolver.Reflection +open import Cubical.Tactics.WildCatSolver.Reflection open import Cubical.Tactics.Reflection open import Agda.Builtin.String From 8d8e5dcba0536ae0496dfe0cdbaa1e80fb167c64 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 16:35:50 +0100 Subject: [PATCH 05/20] cleanup --- Cubical/Tactics/GroupSolver/Example.agda | 8 +- Cubical/Tactics/GroupSolver/Solver.agda | 10 +- Cubical/Tactics/GroupoidSolver/Example.agda | 17 +- Cubical/Tactics/GroupoidSolver/Solver.agda | 4 +- Cubical/Tactics/WildCatSolver/Example.agda | 2 +- Cubical/Tactics/WildCatSolver/Reflection.agda | 20 +- Cubical/Tactics/WildCatSolver/Solver.agda | 4 +- Cubical/Tactics/WildCatSolver/Solvers.agda | 324 +++--------------- Cubical/WildCat/Base.agda | 20 +- Cubical/WildCat/Functor.agda | 6 +- 10 files changed, 96 insertions(+), 319 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 44c4aa3007..95dd69dd79 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -46,16 +46,16 @@ module example (G G* G○ : Group ℓ) pA≡pB : pA ≡ pB pA≡pB = solveGroup (G ∷ []) - - - + + + module T2 p p' q r s t u where lhs rhs : fst G lhs = (p · p') · (inv p' · (F* ⟪ (((*.inv q) *.· r) *.· F○ ⟪ ○.inv t ⟫ *.· (*.inv (F○ ⟪ s ○.· s ⟫) *.· F○ ⟪ u ⟫ )) ⟫ )) - + rhs = inv (F* ⟪ q ⟫ · inv p) · (F* ⟪ r ⟫) · (comp-WildFunctor F○ F*) ⟪ ○.inv (s ○.· s ○.· t) ○.· u ⟫ diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index 2233ef47b7..786084ee42 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -54,14 +54,14 @@ module _ {ℓ} where WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR - WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) + WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f GroupHom' : (G H : Group ℓ) → Type ℓ - + GroupHom' G H = WildFunctor (WildGroupoid.wildCat (Group→WildGroupoid G)) (WildGroupoid.wildCat (Group→WildGroupoid H)) @@ -81,7 +81,7 @@ module _ {ℓ} where (GroupStr.1g (snd H)) (hom1g (G .snd) (WildFunctor.F-hom wf) (H .snd) (WildFunctor.F-seq wf)) - (WildFunctor.F-id wf) i + (WildFunctor.F-id wf) i WildFunctor.F-seq (Iso.leftInv IsoGroupHom' wf i) = λ f g → WildFunctor.F-seq wf f g @@ -92,10 +92,10 @@ module Group-Solver ℓ where WildStr.toWildCat GroupWS = WildGroupoid.wildCat ∘ Group→WildGroupoid WildStr.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) - private + private module GRP-WS = WildStr GroupWS macro solveGroup : R.Term → R.Term → R.TC Unit solveGroup = GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) - + diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 583c4cfe02..34951f6952 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -15,15 +15,12 @@ private variable ℓ ℓ' : Level - - module example (WG WG* : WildGroupoid ℓ ℓ') (F : WildFunctor (WildGroupoid.wildCat WG*) (WildGroupoid.wildCat WG)) where - open WildGroupoid-Solver ℓ ℓ' @@ -45,14 +42,14 @@ module example (WG WG* : WildGroupoid ℓ ℓ') pA pB : Hom[ x , w ] - pA = ((((((id ⋆ p') ⋆ q) ⋆ (inv q)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r - pB = ((id ⋆ p) ⋆ q) ⋆ r + pA = ((((((id ⋆ p') ⋆ q) ⋆ ((inv q) ⋆ id)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r + pB = p ⋆ (q ⋆ r) pA≡pB : pA ≡ pB pA≡pB = solveWildGroupoid (WG ∷ WG* ∷ []) - - - + + + module T2 x y z w (p : Hom[ x , F ⟅ y ⟆ ]) (p' : Hom[ F ⟅ y ⟆ , x ]) @@ -62,7 +59,7 @@ module example (WG WG* : WildGroupoid ℓ ℓ') lhs rhs : Hom[ x , F-ob F w ] lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) - rhs = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) + rhs = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) lhs≡rhs : lhs ≡ rhs @@ -74,7 +71,7 @@ module example (WG WG* : WildGroupoid ℓ ℓ') where lhs rhs : Hom[ obs 1 , obs 5 ] - lhs = homs _ 2 1 ⋆ (homs _ _ 2 ⋆ (homs 3 5 2 ⋆ (homs _ 6 3 ⋆ inv (homs _ 6 3)))) + lhs = homs _ 2 1 ⋆ (homs _ _ 2 ⋆ (homs 3 5 2 ⋆ (homs _ 6 3 ⋆ inv (homs _ 6 3)))) rhs = ((homs _ 2 1 ⋆ homs _ _ 2) ⋆ id) ⋆ homs 3 5 2 diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index 727eff5981..9e9fbe2249 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -47,10 +47,10 @@ module WildGroupoid-Solver ℓ ℓ' where WildStr.toWildCat GroupoidWS = WildGroupoid.wildCat WildStr.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid - private + private module GPD-WS = WildStr GroupoidWS macro solveWildGroupoid : R.Term → R.Term → R.TC Unit solveWildGroupoid = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) - + diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index e982a978a4..edc092f496 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -35,7 +35,7 @@ module exampleC (WC WC* : WildCat ℓ ℓ') open WildCat-Solver ℓ ℓ' pA=pB : pA ≡ pB - pA=pB = solveWildCat (WC ∷ WC* ∷ []) + pA=pB = solveWildCat (WC ∷ WC* ∷ []) pB=pC : pB ≡ pC pB=pC = solveWildCat (WC ∷ WC* ∷ []) diff --git a/Cubical/Tactics/WildCatSolver/Reflection.agda b/Cubical/Tactics/WildCatSolver/Reflection.agda index f0b2c24468..414bdfc61b 100644 --- a/Cubical/Tactics/WildCatSolver/Reflection.agda +++ b/Cubical/Tactics/WildCatSolver/Reflection.agda @@ -26,7 +26,7 @@ x ≟ℕ x₁ = Dec→Bool (discreteℕ x x₁) _<>_ = primStringAppend -_≟qn_ = R.primQNameEquality +_≟qn_ = R.primQNameEquality quotedMaybe→maybeTerm : R.Term → R.TC (Maybe (R.Term)) quotedMaybe→maybeTerm (R.con (quote nothing) _) = R.returnTC nothing @@ -61,7 +61,7 @@ blockIfContainsMeta (R.lit l) = R.returnTC _ blockIfContainsMeta (R.meta m _) = R.blockTC (R.blockerMeta m) blockIfContainsMeta R.unknown = R.returnTC _ -any? : List Bool → Bool +any? : List Bool → Bool any? [] = false any? (false ∷ x₁) = any? x₁ any? (true ∷ x₁) = true @@ -74,7 +74,7 @@ containsMeta? (R.con c args) = containsMetaAny? args containsMeta? (R.def f args) = containsMetaAny? args containsMeta? (R.lam v₁ (R.abs _ t)) = containsMeta? t containsMeta? (R.pat-lam cs args) = containsMetaAny? args -containsMeta? (R.pi (R.arg _ a) (R.abs _ b)) = containsMeta? a or containsMeta? b +containsMeta? (R.pi (R.arg _ a) (R.abs _ b)) = containsMeta? a or containsMeta? b containsMeta? (R.agda-sort s) = false containsMeta? (R.lit l) = false containsMeta? (R.meta x x₁) = true @@ -200,15 +200,15 @@ instance _∷nl_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart -_∷nl_ x y = x ∷ₑ "\n" ∷ₑ y +_∷nl_ x y = x ∷ₑ "\n" ∷ₑ y niceAtomList : List (R.Term) → List R.ErrorPart -niceAtomList = h 0 +niceAtomList = h 0 where - h : _ → _ + h : _ → _ h _ [] = [] - h k (x ∷ xs) = (mkNiceVar k <> " = ") ∷ₑ x ∷ₑ "\n" ∷ₑ h (suc k) xs + h k (x ∷ xs) = (mkNiceVar k <> " = ") ∷ₑ x ∷ₑ "\n" ∷ₑ h (suc k) xs foldlM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} @@ -222,7 +222,7 @@ R∙ x y = R.def (quote _∙_) (x v∷ y v∷ [] ) Rrefl : R.Term Rrefl = R.def (quote refl) [] -unArgs : List (R.Arg (R.Term)) → List R.ErrorPart +unArgs : List (R.Arg (R.Term)) → List R.ErrorPart unArgs [] = [] unArgs (R.arg i x ∷ x₁) = x ∷ₑ unArgs x₁ @@ -233,7 +233,7 @@ getConTail (R.def f args) = "𝒅𝒆𝒇 " ∷ₑ f ∷ₑ " " ∷ₑ unArgs ar getConTail _ = "other..." ∷ₑ [] tryAllTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → - R.TC B → List A → (A → R.TC B) → R.TC B + R.TC B → List A → (A → R.TC B) → R.TC B tryAllTC fallback [] f = fallback tryAllTC fallback (x ∷ xs) f = f x <|> tryAllTC fallback xs f @@ -242,7 +242,7 @@ tryAllTC fallback (x ∷ xs) f = foldPathTerms : List (Maybe R.Term) → Maybe R.Term foldPathTerms [] = nothing foldPathTerms (nothing ∷ xs) = foldPathTerms xs -foldPathTerms (just x ∷ xs) = +foldPathTerms (just x ∷ xs) = just $ Mb.rec x (λ xs' → R.def (quote _∙_) (x v∷ xs' v∷ [])) (foldPathTerms xs) symPathTerms : List (Maybe R.Term) → List (Maybe R.Term) diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index 85bc64d0e3..5ca361865e 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -47,10 +47,10 @@ module WildCat-Solver ℓ ℓ' where WildStr.toWildCat WildCatWS = idfun _ WildStr.mbIsWildGroupoid WildCatWS = nothing - private + private module WC-WS = WildStr WildCatWS macro solveWildCat : R.Term → R.Term → R.TC Unit solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) - + diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index e8f7b5f589..e64929a3a6 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -4,12 +4,8 @@ module Cubical.Tactics.WildCatSolver.Solvers where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure open import Cubical.Foundations.Function as Fu open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv open import Cubical.Data.Bool as 𝟚 hiding (_≤_) open import Cubical.Data.Nat as ℕ hiding (_·_) @@ -17,26 +13,17 @@ open import Cubical.Data.Nat.Order.Recursive open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sigma -open import Cubical.Data.Sum as ⊎ open import Cubical.Data.List as Li open import Cubical.Data.Maybe as Mb - -open import Cubical.HITs.Interval - --- open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary - open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R open import Cubical.Tactics.WildCatSolver.Reflection open import Cubical.Tactics.Reflection open import Agda.Builtin.String --- open import Cubical.WildCat.WGE open import Cubical.WildCat.Base open import Cubical.WildCat.Functor -open import Cubical.Algebra.Group private variable @@ -50,16 +37,15 @@ mvFlagElim (just x) _ = x record TypeWithRel ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where no-eta-equality field - Carrier : Type ℓ + Carrier : Type ℓ _[_∼_] : Carrier → Carrier → Type ℓ' - -- mbIsWildGroupoid : Maybe (∀ WS → IsWildGroupoid (toWildCat WS)) record TypeWithRelMor (T T' : TypeWithRel ℓ ℓ') : Type (ℓ-max ℓ ℓ') where open TypeWithRel field - TF-ob : TypeWithRel.Carrier T → TypeWithRel.Carrier T' + TF-ob : TypeWithRel.Carrier T → TypeWithRel.Carrier T' TF-hom : ∀ {x y} → T [ x ∼ y ] → T' [ TF-ob x ∼ TF-ob y ] @@ -78,10 +64,10 @@ module FuExpr' (ℓ ℓ' : Level) (InvFlag : Type) (𝑭-ob-map : ∀ {T T'} → 𝑭 T T' → TypeWithRel.Carrier (𝑻→twr T) → TypeWithRel.Carrier (𝑻→twr T') ) where - + module _ 𝑻 where open TypeWithRel (𝑻→twr 𝑻) public - + data FuExpr : (T : 𝑻) → Carrier T → Carrier T → Type (ℓ-suc (ℓ-max ℓ ℓ')) where 𝒂⟦_⟧ : ∀ {T x y} → T [ x ∼ y ] → FuExpr T x y idFE : ∀ {T x} → FuExpr T x x @@ -91,7 +77,7 @@ module FuExpr' (ℓ ℓ' : Level) (InvFlag : Type) invFE : ∀ {T x y} (invF : InvFlag) → FuExpr T x y → FuExpr T y x - ⟅_,_,_⟆FE : ∀ T {T' x y} + ⟅_,_,_⟆FE : ∀ T {T' x y} (F : 𝑭 T T') → FuExpr T x y → FuExpr T' (𝑭-ob-map F x) (𝑭-ob-map F y) @@ -106,7 +92,7 @@ module FuExpr' (ℓ ℓ' : Level) (InvFlag : Type) printFuExpr (invFE invF x) = "(" <> printFuExpr x <> ")⁻¹" printFuExpr ⟅ T , F , x ⟆FE = - "⟪" <> printFuExpr x <> "⟫" + "⟪" <> printFuExpr x <> "⟫" module _ InvFlag where module TermExpr = FuExpr' ℓ-zero ℓ-zero InvFlag @@ -128,12 +114,12 @@ module _ if ([w] [F] [a] : List R.Term) where lkW = lookupWithDefault R.unknown [w] lkF = lookupWithDefault R.unknown [F] lkA = lookupWithDefault R.unknown [a] - + ℕExpr→TermExrp : ∀ {w} → ℕE if (lift w) _ _ → TE if (lift (lkW w)) _ _ ℕExpr→TermExrp 𝒂⟦ x ⟧ = 𝒂⟦ lkA x ⟧ ℕExpr→TermExrp idFE = idFE ℕExpr→TermExrp (x ⋆FE x₁) = - (ℕExpr→TermExrp x ⋆FE ℕExpr→TermExrp x₁) + (ℕExpr→TermExrp x ⋆FE ℕExpr→TermExrp x₁) ℕExpr→TermExrp (invFE invF x) = invFE invF (ℕExpr→TermExrp x) ℕExpr→TermExrp (⟅ _ , F , x ⟆FE) = ⟅ _ , lkF F , ℕExpr→TermExrp x ⟆FE @@ -141,7 +127,7 @@ module _ if ([w] [F] [a] : List R.Term) where wc→twr : WildCat ℓ ℓ' → TypeWithRel ℓ ℓ' TypeWithRel.Carrier (wc→twr x) = WildCat.ob x -TypeWithRel._[_∼_] (wc→twr x) = WildCat.Hom[_,_] x +TypeWithRel._[_∼_] (wc→twr x) = WildCat.Hom[_,_] x mbFunctorApp : R.Term → Maybe (R.Term × R.Term) mbFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t @@ -158,9 +144,9 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where ws→twr : wildStr → TypeWithRel ℓ ℓ' ws→twr = wc→twr ∘ toWildCat - + module _ (ws : wildStr) where - open WildCat (toWildCat ws) renaming (Hom[_,_] to _H[_,_] ; _⋆_ to _⟨_⋆_⟩) public + open WildCat (toWildCat ws) renaming (Hom[_,_] to _H[_,_] ; _⋆_ to _⟨_⋆_⟩) public module _ (invF : InvFlag) where WG = (record { wildCat = toWildCat ws ; isWildGroupoid = mvFlagElim mbIsWildGroupoid invF ws }) @@ -168,7 +154,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where WildF : wildStr → wildStr → Type (ℓ-max ℓ ℓ') - WildF ws ws' = WildFunctor (toWildCat ws) (toWildCat ws') + WildF ws ws' = WildFunctor (toWildCat ws) (toWildCat ws') open WildFunctor @@ -183,23 +169,23 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where invFE : ∀ {ws x y} (invF : InvFlag) → (f : ws H[ x , y ]) → FuCases ws {y} {x} (inv ws invF f) - ⟅_,_,_⟆FE : ∀ ws {ws' x y} + ⟅_,_,_⟆FE : ∀ ws {ws' x y} (F : WildF ws ws') → (f : ws H[ x , y ]) → FuCases ws' {F-ob F x} {F-ob F y} (F-hom F f ) module _ where open FuExpr' ℓ ℓ' InvFlag wildStr ws→twr - WildF WildFunctor.F-ob public + WildF WildFunctor.F-ob public evFuExpr : {ws : wildStr} {x y : ob ws} - → FuExpr ws x y → ws H[ x , y ] + → FuExpr ws x y → ws H[ x , y ] evFuExpr FuExpr'.𝒂⟦ x ⟧ = x evFuExpr {ws} FuExpr'.idFE = id ws evFuExpr {ws} (x FuExpr'.⋆FE x₁) = ws ⟨ evFuExpr x ⋆ evFuExpr x₁ ⟩ evFuExpr {ws} (FuExpr'.invFE invF x) = inv ws invF (evFuExpr x) evFuExpr FuExpr'.⟅ T , F , x ⟆FE = F ⟪ evFuExpr x ⟫ - + module _ {ws : wildStr} where data FuAtom : ob ws → ob ws → Type (ℓ-max ℓ ℓ') where @@ -219,12 +205,12 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where invFuAtom x a[ x₁ ] = a[ x , x₁ ]⁻ invFuAtom x a[ x₁ , x₂ ]⁻ = a[ x₂ ] - + ff↓ : ∀ {x y} → FuFlat x y → FuFlat x y ff↓ [ff] = [ff] ff↓ (x ff∷ x₁) = ff↓ x ff∷ x₁ ff↓ (x invol invF ∷ff x₁) = ff↓ x ff∷ x₁ ff∷ invFuAtom invF x₁ - + ff↑ : ∀ {x y} → FuFlat x y → FuFlat x y ff↑ [ff] = [ff] ff↑ (x ff∷ x₁) = ff↑ x ff∷ x₁ @@ -239,7 +225,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where ffInv : ∀ {x y} → InvFlag → FuFlat x y → FuFlat y x ffInv x [ff] = [ff] - ffInv x (x₁ ff∷ x₂) = ([ff] ff∷ (invFuAtom x x₂)) ff++ ffInv x x₁ + ffInv x (x₁ ff∷ x₂) = ([ff] ff∷ (invFuAtom x x₂)) ff++ ffInv x x₁ ffInv x (x₁ invol invF ∷ff x₂) = ([ff] invol invF ∷ff x₂) ff++ ffInv x x₁ @@ -247,7 +233,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where InvFlag → FuAtom y x → FuAtom x y invFuAtomExplicit ws = invFuAtom {ws} - aa⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} + aa⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} → (F : WildF ws' ws) → (FuAtom x y) → FuAtom (F-ob F x) (F-ob F y) @@ -255,7 +241,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where aa⟪ F , a[ x , x₁ ]⁻ ⟫ = a[ x , F ⟪ x₁ ⟫ ]⁻ - ff⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} + ff⟪_,_⟫ : ∀ {ws' ws : wildStr} {x y} → (F : WildF ws' ws) → (FuFlat x y) → FuFlat (F-ob F x) (F-ob F y) @@ -265,7 +251,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where FuExpr→FF : {ws : wildStr} {x y : ob ws} - → FuExpr ws x y → FuFlat x y + → FuExpr ws x y → FuFlat x y FuExpr→FF 𝒂⟦ x ⟧ = [ff] ff∷ a[ x ] FuExpr→FF idFE = [ff] FuExpr→FF (x ⋆FE x₁) = (FuExpr→FF x) ff++ (FuExpr→FF x₁) @@ -273,20 +259,20 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where FuExpr→FF ⟅ T , F , x ⟆FE = ff⟪ F , FuExpr→FF x ⟫ evAtom : {ws : wildStr} {x y : ob ws} - → FuAtom x y → ws H[ x , y ] + → FuAtom x y → ws H[ x , y ] evAtom a[ x ] = x evAtom {ws} a[ x , x₁ ]⁻ = inv ws x x₁ invFuAtom-involR : ∀ ws {x y} invF → (a : FuAtom {ws} y x) → - (ws ⟨ evAtom a ⋆ evAtom (invFuAtom invF a) ⟩) ≡ id ws + (ws ⟨ evAtom a ⋆ evAtom (invFuAtom invF a) ⟩) ≡ id ws invFuAtom-involR ws invF a[ x ] = ⋆InvR ws _ _ invFuAtom-involR ws invF a[ x , x₁ ]⁻ = ⋆InvL ws _ _ evFF : {ws : wildStr} {x y : ob ws} - → FuFlat x y → ws H[ x , y ] + → FuFlat x y → ws H[ x , y ] evFF {ws} [ff] = id ws evFF {ws} (x ff∷ x₁) = ws ⟨ (evFF x) ⋆ (evAtom x₁) ⟩ evFF (x invol invF ∷ff x₁) = evFF x @@ -302,9 +288,9 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where ⋆Assoc ws _ _ _ ∙∙ cong₂ (ws ⟨_⋆_⟩) (evFF≡↓ ws f) (invFuAtom-involR ws invF x) ∙∙ ⋆IdR ws (evFF f) - + evFF++ : ∀ {ws x y z} → (g : FuFlat {ws} x y) → (h : FuFlat y z) → - ws ⟨ evFF g ⋆ evFF h ⟩ ≡ (evFF (g ff++ h)) + ws ⟨ evFF g ⋆ evFF h ⟩ ≡ (evFF (g ff++ h)) evFF++ {ws = ws} g [ff] = ⋆IdR ws _ evFF++ {ws = ws} g (h ff∷ x) = sym (⋆Assoc ws _ _ _) ∙ cong (ws ⟨_⋆ (evAtom x) ⟩) (evFF++ g h) @@ -314,7 +300,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where (g : FuAtom {ws} x y) → inv ws invF (evAtom g) ≡ evAtom (invFuAtom invF g) evAinv invF a[ x ] = refl - evAinv {ws = ws} invF a[ x , x₁ ]⁻ with mbIsWildGroupoid | invol-inv ws + evAinv {ws = ws} invF a[ x , x₁ ]⁻ with mbIsWildGroupoid | invol-inv ws ... | just x₂ | ii = ii _ x₁ @@ -332,15 +318,15 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where aa-Func-inv invF F a[ x ] = refl aa-Func-inv {ws' = ws'} invF F a[ x , x₁ ]⁻ with mbIsWildGroupoid | invol-inv ws' ... | just x₂ | ii = ii _ _ - + evFFinv : ∀ {ws x y} → (invF : InvFlag) → (g : FuFlat {ws} x y) → - inv ws invF (evFF g) ≡ evFF (ffInv invF g) + inv ws invF (evFF g) ≡ evFF (ffInv invF g) evFFinv {ws} invF [ff] = id≡inv-id ws invF evFFinv {ws} invF (g ff∷ x) = distInv ws invF _ _ ∙∙ cong₂ (ws ⟨_⋆_⟩) (evAinv invF x ∙ sym (⋆IdL ws _)) (evFFinv invF g) - ∙∙ evFF++ _ (ffInv invF g) + ∙∙ evFF++ _ (ffInv invF g) evFFinv {ws} invF (g invol invF₁ ∷ff x) = evFFinv invF g ∙∙ sym (⋆IdL ws _) ∙∙ evFF++ _ (ffInv invF g) @@ -356,13 +342,13 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where ff⟪⟫inv : ∀ {ws ws' x y} invF (F : WildFunctor (toWildCat ws) (toWildCat ws')) → (f : FuFlat {ws} x y) → - inv ws' invF (evFF ff⟪ F , f ⟫) + inv ws' invF (evFF ff⟪ F , f ⟫) ≡ evFF ff⟪ F , ffInv invF f ⟫ ff⟪⟫inv {ws' = ws'} invF F [ff] = id≡inv-id ws' invF ff⟪⟫inv {ws' = ws'} invF F (f ff∷ x) = distInv ws' invF _ _ ∙∙ cong₂ (ws' ⟨_⋆_⟩) (aa-Func-inv invF F x ∙ sym (⋆IdL ws' _)) (ff⟪⟫inv invF F f) - ∙∙ ff⟪⟫⋆ F _ (ffInv invF f) + ∙∙ ff⟪⟫⋆ F _ (ffInv invF f) ff⟪⟫inv {ws' = ws'} invF F (f invol invF₁ ∷ff x) = ff⟪⟫inv invF F f ∙∙ sym (⋆IdL ws' _) @@ -377,8 +363,8 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where ff-Func {ws' = ws'} F (f ff∷ x) = F-seq F _ _ ∙ cong₂ (ws' ⟨_⋆_⟩) (ff-Func F f) (aa-Func F x) ff-Func F (f invol invF ∷ff x) = ff-Func F f - - evFF-Func : ∀ {ws ws'} (F : WildFunctor (toWildCat ws) (toWildCat ws')) {x y} → + + evFF-Func : ∀ {ws ws'} (F : WildFunctor (toWildCat ws) (toWildCat ws')) {x y} → (f : FuExpr ws x y) → F-hom F (evFuExpr f) ≡ evFF {ws'} ff⟪ F , FuExpr→FF f ⟫ evFF-Func {ws' = ws'} F FuExpr'.𝒂⟦ x ⟧ = sym (⋆IdL ws' _) @@ -393,7 +379,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where evFF-Func F FuExpr'.⟅ T , F' , f ⟆FE = cong (F-hom F) (evFF-Func F' f) ∙ ff-Func F ff⟪ F' , FuExpr→FF f ⟫ - + evFuExpr≡evFF : (ws : wildStr) {x y : ob ws} → (f : FuExpr ws x y) → evFuExpr f ≡ evFF (FuExpr→FF f) @@ -402,27 +388,19 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where evFuExpr≡evFF ws (f FuExpr'.⋆FE f₁) = cong₂ (ws ⟨_⋆_⟩) (evFuExpr≡evFF ws f) (evFuExpr≡evFF ws f₁) ∙ evFF++ (FuExpr→FF f) (FuExpr→FF f₁) - evFuExpr≡evFF ws (FuExpr'.invFE invF f) = - cong (inv ws invF) (evFuExpr≡evFF ws f) ∙ evFFinv invF (FuExpr→FF f) + evFuExpr≡evFF ws (FuExpr'.invFE invF f) = + cong (inv ws invF) (evFuExpr≡evFF ws f) ∙ evFFinv invF (FuExpr→FF f) evFuExpr≡evFF ws FuExpr'.⟅ T , F , f ⟆FE = evFF-Func F f - -- evFuExpr-singl : {ws : wildStr} {x y : ob ws} - -- → (f : FuExpr ws x y) → - -- Σ {!!} {!λ f' → evFuExpr f ≡ evFF (FuExpr→FF f)!} - -- evFuExpr-singl = {!!} - - - magicNumber : ℕ magicNumber = 100 infixl 5 us∷_ - -- us∷_ : R.Term us∷_ : List (R.Arg R.Term) → List (R.Arg R.Term) us∷_ = R.unknown v∷_ - - buildFromTE : ∀ {W} → TE InvFlag (lift W) _ _ → R.Term + + buildFromTE : ∀ {W} → TE InvFlag (lift W) _ _ → R.Term buildFromTE FuExpr'.𝒂⟦ x ⟧ = R.con (quote FuExpr'.𝒂⟦_⟧) ([ varg x ]) buildFromTE FuExpr'.idFE = R.con (quote FuExpr'.idFE) [] buildFromTE (x FuExpr'.⋆FE x₁) = @@ -434,7 +412,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where buildFromTE FuExpr'.⟅ lift T , F , x ⟆FE = R.con (quote FuExpr'.⟅_,_,_⟆FE) (T v∷ F v∷ buildFromTE x v∷ []) - + module tryWCE WS (tGs : List R.Term) where @@ -471,50 +449,25 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where where h : ℕ → R.Term → R.TC (List R.Term) h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) - h (suc k) t = + h (suc k) t = (mb-invol' W t) >>= Mb.rec (R.returnTC []) - λ (t' , t'↓) → do + λ (t' , t'↓) → do p' ← h k t'↓ - -- R.typeError [ R.termErr t' ] R.returnTC (t' ∷ p') - + redList' : R.Term → R.Term → R.TC (List R.Term) redList' W = redList W magicNumber - -- involPathTerm : R.Term → ℕ → R.Term → R.TC R.Term - -- involPathTerm W (suc n) t = - -- (mb-invol' W t) >>= - -- Mb.rec - -- (R.returnTC (R.def (quote refl) [])) - -- λ (t' , t'↓) → do - -- p' ← involPathTerm W n t'↓ - -- R.returnTC $ R.def (quote _∙_) - -- ((R.def (quote evFF≡↓) (WS v∷ W v∷ t' v∷ [])) v∷ p' v∷ []) - -- involPathTerm _ zero _ = - -- R.typeError ("magic number too low in mb-invol" ∷ₑ []) - - -- involPathTerm' : R.Term → R.Term → R.TC R.Term - -- involPathTerm' W t = involPathTerm W magicNumber t - -- -- R.checkType pa (R.def (quote Path) {!!}) - checkFromTE : ∀ {W} → TE InvFlag (lift W) _ _ → - R.TC R.Term --(Σ wildStr λ ws → Σ _ (λ x → Σ _ (FuExpr ws x))) + R.TC R.Term checkFromTE {W} te = do let te' = buildFromTE te R.checkType te' (R.def (quote FuExpr) (WS v∷ W v∷ us∷ us∷ [] )) - -- (R.con (quote _,_) - -- (W - -- v∷ R.con (quote _,_) - -- (R.unknown v∷ R.con (quote _,_) - -- (R.unknown v∷ te' v∷ []) v∷ []) v∷ []) ) - -- ? - - -- if = {!!} - + tryE : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) fromWC : R.Term → R.TC R.Term @@ -587,7 +540,7 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where h : R.Term → R.TC (R.Term × R.Term) h (R.con _ l) = match2Vargs l - h (R.def _ l) = match2Vargs l + h (R.def _ l) = match2Vargs l h t = R.typeError $ "match2Fail: " ∷ₑ t ∷ₑ [] @@ -609,10 +562,8 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where R.catchTC (tryId W t) (R.catchTC (tryInv W k t) - -- (tryOp k t) (R.catchTC (tryOp W k t) - (R.catchTC (tryFunc W k t) (atom W t))) - ) + (R.catchTC (tryFunc W k t) (atom W t)))) @@ -630,8 +581,8 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 expr0 ← tryWCE.checkFromTE Ws Wts t0' - expr1 ← tryWCE.checkFromTE Ws Wts t1' - -- R.typeError + expr1 ← tryWCE.checkFromTE Ws Wts t1' + let msg = (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) invol0 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr0 ])) @@ -647,187 +598,16 @@ record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where invPa1 = Li.map (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) red1 - -- R.typeError ( - -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red0)) - -- ++ₑ " " ∷nl " " ∷nl - -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red1)) ++ₑ []) - --- -- invPa0 ← tryWCE.involPathTerm' Ws Wts Wt invol0 -->>= R.normalise --- -- invPa1 ← tryWCE.involPathTerm' Ws Wts Wt invol1 -->>= R.normalise - --- -- -- R.typeError (invol0 ∷nl --- -- -- invol1 ∷ₑ []) - let finalTrm0 = just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) ∷ invPa0 - -- R.def (quote _∙_) - -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) v∷ - -- invPa0 v∷ []) + finalTrm1 = just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) ∷ invPa1 - -- R.def (quote _∙_) - -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) v∷ - -- invPa1 v∷ []) let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms (finalTrm0 ++ symPathTerms finalTrm1) - -- --invPa0 - -- R.def (quote _∙_) - -- (finalTrm0 v∷ --R.unknown v∷ - -- (R.def (quote sym) - -- (finalTrm1 v∷ [])) v∷ []) - - -- msg ← R.formatErrorParts (finalTrm ∷ₑ []) - - -- R.typeError (msg ∷ₑ []) - - (R.unify finalTrm hole) -- <|> R.typeError msg - - --- macro - - --- testInferExpr : R.Term → R.Term → R.Term → R.Term → R.TC Unit --- testInferExpr Ws Wts' t ho = do --- Wts ← quotedList→ListOfTerms Wts' --- Wt ← tryAllTC --- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) --- Wts R.returnTC --- expr ← tryWCE.tryE Ws Wts Wt magicNumber t -->> R.returnTC _ - --- tryWCE.checkFromTE Ws Wts expr --- R.typeError (TermExpr.printFuExpr InvFlag (λ _ → "●") expr ∷ₑ []) --- -- R.unify R.unknown ho - - --- -- solveNoInvs : R.Term → R.Term → R.Term → R.TC Unit --- -- solveNoInvs Ws Wts' hole = do --- -- Wts ← quotedList→ListOfTerms Wts' --- -- Wt ← tryAllTC --- -- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) --- -- Wts R.returnTC --- -- (t0 , t1) ← R.withNormalisation true $ --- -- R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec --- -- (R.typeError [ R.strErr "unable to get boundary" ]) --- -- (λ x → R.returnTC x) --- -- t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 --- -- t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 --- -- expr0 ← tryWCE.checkFromTE Ws Wts t0' --- -- expr1 ← tryWCE.checkFromTE Ws Wts t1' --- -- -- R.typeError --- -- -- R.typeError (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl --- -- -- TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) - --- -- let finalTrm = R.def (quote _∙∙_∙∙_) --- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ expr0 v∷ [])) v∷ --- -- (R.def (quote refl) []) v∷ --- -- (R.def (quote sym) --- -- ((R.def (quote evFuExpr≡evFF)) (Ws v∷ expr1 v∷ []) v∷ [])) v∷ []) --- -- R.unify finalTrm hole - - --- solveW : R.Term → R.Term → R.Term → R.TC Unit --- solveW Ws Wts' hole = do --- Wts ← quotedList→ListOfTerms Wts' --- Wt ← tryAllTC --- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) --- Wts R.returnTC --- hoTy ← R.withNormalisation true $ --- R.inferType hole >>= wait-for-type --- (t0 , t1) ← (get-boundary hoTy ) >>= Mb.rec --- (R.typeError [ R.strErr "unable to get boundary" ]) --- (λ x → R.returnTC x) --- t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 --- t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 --- expr0 ← tryWCE.checkFromTE Ws Wts t0' --- expr1 ← tryWCE.checkFromTE Ws Wts t1' --- -- R.typeError --- let msg = (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl --- TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) --- invol0 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr0 ])) --- invol1 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr1 ])) - - --- red0 ← tryWCE.redList' Ws Wts Wt invol0 --- red1 ← tryWCE.redList' Ws Wts Wt invol1 - --- let invPa0 = Li.map --- (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) --- red0 --- invPa1 = Li.map --- (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) --- red1 --- -- R.typeError ( --- -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red0)) --- -- ++ₑ " " ∷nl " " ∷nl --- -- (join (Li.map (λ x → "\n-- " ∷ₑ [ R.termErr x ] ) red1)) ++ₑ []) - --- -- -- invPa0 ← tryWCE.involPathTerm' Ws Wts Wt invol0 -->>= R.normalise --- -- -- invPa1 ← tryWCE.involPathTerm' Ws Wts Wt invol1 -->>= R.normalise - --- -- -- -- R.typeError (invol0 ∷nl --- -- -- -- invol1 ∷ₑ []) - - - --- let finalTrm0 = --- just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) --- ∷ invPa0 --- -- R.def (quote _∙_) --- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr0 v∷ [])) v∷ --- -- invPa0 v∷ []) --- finalTrm1 = --- just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) --- ∷ invPa1 --- -- R.def (quote _∙_) --- -- ((R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) v∷ --- -- invPa1 v∷ []) - --- let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms --- (finalTrm0 ++ symPathTerms finalTrm1) --- -- --invPa0 --- -- R.def (quote _∙_) --- -- (finalTrm0 v∷ --R.unknown v∷ --- -- (R.def (quote sym) --- -- (finalTrm1 v∷ [])) v∷ []) - --- -- msg ← R.formatErrorParts (finalTrm ∷ₑ []) - --- -- R.typeError (msg ∷ₑ []) - --- (R.unify finalTrm hole) -- <|> R.typeError msg - - - --- module _ (G : Group ℓ) where --- open GroupStr (snd G) --- Group→WildGroupoid : WildGroupoid ℓ-zero ℓ --- WildCat.ob (WildGroupoid.wildCat Group→WildGroupoid) = Unit --- WildCat.Hom[_,_] (WildGroupoid.wildCat Group→WildGroupoid) _ _ = ⟨ G ⟩ --- WildCat.id (WildGroupoid.wildCat Group→WildGroupoid) = 1g --- WildCat._⋆_ (WildGroupoid.wildCat Group→WildGroupoid) = _·_ --- WildCat.⋆IdL (WildGroupoid.wildCat Group→WildGroupoid) = ·IdL --- WildCat.⋆IdR (WildGroupoid.wildCat Group→WildGroupoid) = ·IdR --- WildCat.⋆Assoc (WildGroupoid.wildCat Group→WildGroupoid) _ _ _ = sym (·Assoc _ _ _) --- wildIsIso.inv' (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = inv f --- wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f --- wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f - - --- GroupWS : WildStr ℓ-zero ℓ --- GroupoidWS WildCatWS : WildStr ℓ ℓ' - --- WildStr.wildStr (GroupWS {ℓ}) = Group ℓ --- WildStr.toWildCat (GroupWS) = WildGroupoid.wildCat ∘ Group→WildGroupoid --- WildStr.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) - --- WildStr.wildStr (GroupoidWS {ℓ} {ℓ'}) = WildGroupoid ℓ ℓ' --- WildStr.toWildCat GroupoidWS = WildGroupoid.wildCat --- WildStr.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid - - - + (R.unify finalTrm hole) diff --git a/Cubical/WildCat/Base.agda b/Cubical/WildCat/Base.agda index fbaa83fc3a..c9ea3e5428 100644 --- a/Cubical/WildCat/Base.agda +++ b/Cubical/WildCat/Base.agda @@ -88,31 +88,31 @@ _^op : WildCat ℓ ℓ' → WildCat ℓ ℓ' (C ^op) .⋆IdR = C .⋆IdL (C ^op) .⋆Assoc f g h = sym (C .⋆Assoc _ _ _) -IsWildGroupoid : (wc : WildCat ℓ ℓ') → Type (ℓ-max ℓ ℓ') +IsWildGroupoid : (wc : WildCat ℓ ℓ') → Type (ℓ-max ℓ ℓ') IsWildGroupoid wc = ∀ {x y} f → wildIsIso {C = wc} {x} {y} f - - + + record WildGroupoid ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where no-eta-equality - + field - wildCat : WildCat ℓ ℓ' + wildCat : WildCat ℓ ℓ' isWildGroupoid : IsWildGroupoid wildCat - _⋆'_ = _⋆_ wildCat + _⋆'_ = _⋆_ wildCat inv : ∀ {x y} → wildCat [ x , y ] → wildCat [ y , x ] inv f = wildIsIso.inv' (isWildGroupoid f) ⋆InvR : ∀ {x y} → (f : wildCat [ x , y ]) → f ⋆' inv f ≡ id wildCat ⋆InvR f = wildIsIso.retr (isWildGroupoid f) - + ⋆InvL : ∀ {x y} → (f : wildCat [ x , y ]) → inv f ⋆' f ≡ id wildCat ⋆InvL f = wildIsIso.sect (isWildGroupoid f) invol-inv : ∀ {x y} → (f : wildCat [ x , y ]) → inv (inv f) ≡ f - invol-inv f = + invol-inv f = (sym (⋆IdL wildCat (inv (inv f))) ∙ cong (_⋆' (inv (inv f))) (sym (⋆InvR _)) ) ∙ ⋆Assoc wildCat _ _ _ ∙ cong (f ⋆'_) (⋆InvR (inv f)) ∙ ⋆IdR wildCat f @@ -124,7 +124,7 @@ record WildGroupoid ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where invUniqueR : ∀ {x y} {g : wildCat [ x , y ]} {h} → g ⋆' h ≡ id wildCat → h ≡ inv g invUniqueR {g = g} {h} p = (sym (⋆IdL wildCat h) ∙∙ cong (_⋆' h) (sym (⋆InvL g)) - ∙∙ ⋆Assoc wildCat (inv g) g h) ∙∙ cong ((inv g) ⋆'_) p ∙∙ ⋆IdR wildCat (inv g) + ∙∙ ⋆Assoc wildCat (inv g) g h) ∙∙ cong ((inv g) ⋆'_) p ∙∙ ⋆IdR wildCat (inv g) distInv : ∀ {x y z} (f : wildCat [ x , y ]) (g : wildCat [ y , z ]) → @@ -133,7 +133,7 @@ record WildGroupoid ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where (sym (⋆Assoc wildCat _ _ _) ∙∙ cong (_⋆' inv f) (⋆Assoc wildCat _ _ _) ∙∙ (λ i → ⋆Assoc wildCat f ((⋆InvR g) i) (inv f) i)) ∙∙ cong (f ⋆'_) (⋆IdL wildCat _) ∙∙ ⋆InvR f - + id≡inv-id : ∀ {x} → inv (id wildCat) ≡ id wildCat {x = x} id≡inv-id = sym (⋆IdL wildCat (inv (id wildCat))) ∙ ⋆InvR (id wildCat) diff --git a/Cubical/WildCat/Functor.agda b/Cubical/WildCat/Functor.agda index ba579d5d36..1e26f828aa 100644 --- a/Cubical/WildCat/Functor.agda +++ b/Cubical/WildCat/Functor.agda @@ -166,7 +166,7 @@ WildFunctor.F-seq commFunctor _ _ = refl module _ {C : WildGroupoid ℓC ℓC'} {D : WildGroupoid ℓD ℓD'} - (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where + (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where private module gC = WildGroupoid C @@ -177,8 +177,8 @@ module _ module _ (C : WildGroupoid ℓC ℓC') (D : WildGroupoid ℓD ℓD') - (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where - + (F : WildFunctor (WildGroupoid.wildCat C) (WildGroupoid.wildCat D)) where + private module gC = WildGroupoid C module gD = WildGroupoid D From 944dcee557509209e96417e579824dbe3d3dce21 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 16:42:40 +0100 Subject: [PATCH 06/20] cleanup --- Cubical/Tactics/GroupSolver/Example.agda | 2 -- Cubical/Tactics/GroupoidSolver/Example.agda | 6 +++--- 2 files changed, 3 insertions(+), 5 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 95dd69dd79..6c1fa607b7 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -14,8 +14,6 @@ private variable ℓ : Level - - module example (G G* G○ : Group ℓ) (F* : GroupHom' G* G) (F○ : GroupHom' G○ G*) diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 34951f6952..21ed9ab6f9 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -57,9 +57,9 @@ module example (WG WG* : WildGroupoid ℓ ℓ') (r : *.Hom[ z , w ]) where - lhs rhs : Hom[ x , F-ob F w ] - lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F-hom F ((*.inv q) *.⋆ r))) - rhs = inv ((F-hom F q) ⋆ inv p) ⋆ (F-hom F r) + lhs rhs : Hom[ x , F ⟅ w ⟆ ] + lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F ⟪ (*.inv q) *.⋆ r ⟫)) + rhs = inv (F ⟪ q ⟫ ⋆ inv p) ⋆ F ⟪ r ⟫ lhs≡rhs : lhs ≡ rhs From e74b4293283d88395a877631946962edeb2c93d8 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 18:46:06 +0100 Subject: [PATCH 07/20] works --- Cubical/Categories/Category/Base.agda | 16 ++++ Cubical/Tactics/GroupSolver/Solver.agda | 39 +++------ Cubical/Tactics/GroupoidSolver/Example.agda | 36 ++++++++- Cubical/Tactics/GroupoidSolver/Solver.agda | 72 +++++++++-------- Cubical/Tactics/WildCatSolver/Example.agda | 90 ++++++++++++++++----- Cubical/Tactics/WildCatSolver/Solver.agda | 67 +++++++-------- Cubical/Tactics/WildCatSolver/Solvers.agda | 2 +- 7 files changed, 206 insertions(+), 116 deletions(-) diff --git a/Cubical/Categories/Category/Base.agda b/Cubical/Categories/Category/Base.agda index 427b17b94f..854d0718b0 100644 --- a/Cubical/Categories/Category/Base.agda +++ b/Cubical/Categories/Category/Base.agda @@ -161,3 +161,19 @@ isIsoΣPropCat : {C : Category ℓ ℓ'} {P : ℙ (ob C)} inv (isIsoΣPropCat p q f isIsoF) = isIsoF .inv sec (isIsoΣPropCat p q f isIsoF) = isIsoF .sec ret (isIsoΣPropCat p q f isIsoF) = isIsoF .ret + +IsGroupoidCat : (C : Category ℓ ℓ') → Type (ℓ-max ℓ ℓ') +IsGroupoidCat C = ∀ {x} {y} (f : C [ x , y ]) → isIso C f + +IsPropIsGroupoidCat : (C : Category ℓ ℓ') → isProp (IsGroupoidCat C) +IsPropIsGroupoidCat C = + isPropImplicitΠ2 λ _ _ → isPropΠ isPropIsIso + +record GroupoidCat ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + field + category : Category ℓ ℓ' + isGroupoidCat : IsGroupoidCat category + + + _⁻¹ : ∀ {x y} → category [ x , y ] → category [ y , x ] + f ⁻¹ = isIso.inv (isGroupoidCat f) diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index 786084ee42..602abeb41b 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -2,39 +2,20 @@ module Cubical.Tactics.GroupSolver.Solver where - open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function as Fu -open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Foundations.Function open import Cubical.Data.Unit -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sigma -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - --- open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary +open import Cubical.Data.List +open import Cubical.Data.Maybe open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String --- open import Cubical.WildCat.WGE open import Cubical.WildCat.Base +open import Cubical.Tactics.WildCatSolver.Solvers + open import Cubical.WildCat.Functor open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms @@ -87,13 +68,13 @@ module _ {ℓ} where module Group-Solver ℓ where - GroupWS : WildStr ℓ-zero ℓ - WildStr.wildStr GroupWS = Group ℓ - WildStr.toWildCat GroupWS = WildGroupoid.wildCat ∘ Group→WildGroupoid - WildStr.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) + GroupWS : WildCatInstance ℓ-zero ℓ + WildCatInstance.wildStr GroupWS = Group ℓ + WildCatInstance.toWildCat GroupWS = WildGroupoid.wildCat ∘ Group→WildGroupoid + WildCatInstance.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) private - module GRP-WS = WildStr GroupWS + module GRP-WS = WildCatInstance GroupWS macro solveGroup : R.Term → R.Term → R.TC Unit diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 21ed9ab6f9..182996eb9f 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -8,7 +8,7 @@ open import Cubical.WildCat.Base open import Cubical.Tactics.GroupoidSolver.Solver open import Cubical.Data.List open import Cubical.Data.Nat - +open import Cubical.Categories.Category open import Cubical.WildCat.Functor private @@ -77,3 +77,37 @@ module example (WG WG* : WildGroupoid ℓ ℓ') lhs≡rhs : lhs ≡ rhs lhs≡rhs = solveWildGroupoid (WG ∷ []) + + +module exampleGpd (G G* : GroupoidCat ℓ ℓ') + (F : WildFunctor + (WildGroupoid.wildCat (Groupoid-Solver.Groupoid→WildGroupoid _ _ G*)) + (WildGroupoid.wildCat (Groupoid-Solver.Groupoid→WildGroupoid _ _ G))) + where + + open Groupoid-Solver ℓ ℓ' + + + + module * where + open GroupoidCat G* public + open Category category public + + open GroupoidCat G + open Category category + + + module T1 x y z w + (p : Hom[ x , F ⟅ y ⟆ ]) + (p' : Hom[ F ⟅ y ⟆ , x ]) + (q : *.Hom[ z , y ]) + (r : *.Hom[ z , w ]) where + + + lhs rhs : Hom[ x , F ⟅ w ⟆ ] + lhs = (p ⋆ p') ⋆ (p' ⁻¹ ⋆ (F ⟪ (q *.⁻¹) *.⋆ r ⟫)) + rhs = (F ⟪ q ⟫ ⋆ p ⁻¹) ⁻¹ ⋆ F ⟪ r ⟫ + + + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (G ∷ G* ∷ []) diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index 9e9fbe2249..a5cd64d754 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -4,51 +4,57 @@ module Cubical.Tactics.GroupoidSolver.Solver where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function as Fu -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Nat.Order.Recursive +open import Cubical.Foundations.Function open import Cubical.Data.Unit -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sigma -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - --- open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary +open import Cubical.Data.List +open import Cubical.Data.Maybe open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String --- open import Cubical.WildCat.WGE open import Cubical.WildCat.Base -open import Cubical.WildCat.Functor -open import Cubical.Algebra.Group - +open import Cubical.Categories.Category open import Cubical.Tactics.WildCatSolver.Solvers - +open import Cubical.Tactics.WildCatSolver.Solver module WildGroupoid-Solver ℓ ℓ' where - GroupoidWS : WildStr ℓ ℓ' - WildStr.wildStr GroupoidWS = WildGroupoid ℓ ℓ' - WildStr.toWildCat GroupoidWS = WildGroupoid.wildCat - WildStr.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid + GroupoidWS : WildCatInstance ℓ ℓ' + WildCatInstance.wildStr GroupoidWS = WildGroupoid ℓ ℓ' + WildCatInstance.toWildCat GroupoidWS = WildGroupoid.wildCat + WildCatInstance.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid + + private + module WGPD-WS = WildCatInstance GroupoidWS + + macro + solveWildGroupoid : R.Term → R.Term → R.TC Unit + solveWildGroupoid = WGPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) + +module Groupoid-Solver ℓ ℓ' where + + + Groupoid→WildGroupoid : GroupoidCat ℓ ℓ' → WildGroupoid ℓ ℓ' + WildGroupoid.wildCat (Groupoid→WildGroupoid x) = + Cat-Solver.Cat→WildCat _ _ (GroupoidCat.category x) + WildGroupoid.isWildGroupoid (Groupoid→WildGroupoid x) f = wgi + where + + open isIso (GroupoidCat.isGroupoidCat x f) + + wgi : wildIsIso f + wildIsIso.inv' wgi = inv + wildIsIso.sect wgi = sec + wildIsIso.retr wgi = ret + + GroupoidWS : WildCatInstance ℓ ℓ' + WildCatInstance.wildStr GroupoidWS = GroupoidCat ℓ ℓ' + WildCatInstance.toWildCat GroupoidWS = WildGroupoid.wildCat ∘ Groupoid→WildGroupoid + WildCatInstance.mbIsWildGroupoid GroupoidWS = + just (WildGroupoid.isWildGroupoid ∘ Groupoid→WildGroupoid) private - module GPD-WS = WildStr GroupoidWS + module GPD-WS = WildCatInstance GroupoidWS macro solveWildGroupoid : R.Term → R.Term → R.TC Unit diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index edc092f496..1d725db2d3 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -7,35 +7,87 @@ open import Cubical.Foundations.Prelude open import Cubical.WildCat.Base open import Cubical.Tactics.WildCatSolver.Solver open import Cubical.Data.List - open import Cubical.WildCat.Functor private variable ℓ ℓ' : Level -module exampleC (WC WC* : WildCat ℓ ℓ') - (F : WildFunctor WC* WC) where - open WildCat WC - module * = WildCat WC* +module exampleWC where + + module _ (WC WC* : WildCat ℓ ℓ') + (F : WildFunctor WC* WC) where + + open WildCat WC + module * = WildCat WC* + + module _ x y z w v + (p : Hom[ x , F ⟅ y ⟆ ]) + (q : *.Hom[ y , z ]) + (r : *.Hom[ z , w ]) + (s : Hom[ F ⟅ w ⟆ , v ]) where + + + pA pB pC : Hom[ x , v ] + pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) + pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) + pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) + + open WildCat-Solver ℓ ℓ' + + pA=pB : pA ≡ pB + pA=pB = solveWildCat (WC ∷ WC* ∷ []) + + pB=pC : pB ≡ pC + pB=pC = solveWildCat (WC ∷ WC* ∷ []) + + +module exampleC ℓ ℓ' where + open import Cubical.Categories.Category + open Cat-Solver ℓ ℓ' + + + module _ (C C* : Category ℓ ℓ') (F : Functor' C* C) where + + + open Category C + module * = Category C* + + module E1 x y z w v + (p : Hom[ x , y ]) + (q : Hom[ y , z ]) + (r : Hom[ z , w ]) + (s : Hom[ w , v ]) where + + + pA pB pC : Hom[ x , v ] + pA = (p ⋆ (id ⋆ q)) ⋆ (r ⋆ s) + pB = (((p ⋆ q) ⋆ (id ⋆ id) ) ⋆ id ⋆ id) ⋆ (( r ⋆ id) ⋆ s) + pC = p ⋆ (q ⋆ r ⋆ s) + + + pA=pB : pA ≡ pB + pA=pB = solveCat (C ∷ []) + + pB=pC : pB ≡ pC + pB=pC = solveCat (C ∷ []) - module _ x y z w v - (p : Hom[ x , F ⟅ y ⟆ ]) - (q : *.Hom[ y , z ]) - (r : *.Hom[ z , w ]) - (s : Hom[ F ⟅ w ⟆ , v ]) where + module E2 x y z w v + (p : Hom[ x , F ⟅ y ⟆ ]) + (q : *.Hom[ y , z ]) + (r : *.Hom[ z , w ]) + (s : Hom[ F ⟅ w ⟆ , v ]) where - pA pB pC : Hom[ x , v ] - pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) - pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) - pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) + pA pB pC : Hom[ x , v ] + pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) + pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) + pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) - open WildCat-Solver ℓ ℓ' + pA=pB : pA ≡ pB + pA=pB = solveCat (C ∷ C* ∷ []) - pA=pB : pA ≡ pB - pA=pB = solveWildCat (WC ∷ WC* ∷ []) + pB=pC : pB ≡ pC + pB=pC = solveCat (C ∷ C* ∷ []) - pB=pC : pB ≡ pC - pB=pC = solveWildCat (WC ∷ WC* ∷ []) diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index 5ca361865e..5e7b3f84e5 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -2,55 +2,56 @@ module Cubical.Tactics.WildCatSolver.Solver where - open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function as Fu -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Nat.Order.Recursive open import Cubical.Data.Unit -open import Cubical.Data.Empty as ⊥ -open import Cubical.Data.Sigma -open import Cubical.Data.Sum as ⊎ -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - --- open import Cubical.Relation.Nullary -open import Cubical.Relation.Binary +open import Cubical.Data.List +open import Cubical.Data.Maybe open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String --- open import Cubical.WildCat.WGE open import Cubical.WildCat.Base open import Cubical.WildCat.Functor -open import Cubical.Algebra.Group - open import Cubical.Tactics.WildCatSolver.Solvers - +open import Cubical.Categories.Category module WildCat-Solver ℓ ℓ' where - WildCatWS : WildStr ℓ ℓ' - WildStr.wildStr (WildCatWS) = WildCat ℓ ℓ' - WildStr.toWildCat WildCatWS = idfun _ - WildStr.mbIsWildGroupoid WildCatWS = nothing + WildCatWS : WildCatInstance ℓ ℓ' + WildCatInstance.wildStr (WildCatWS) = WildCat ℓ ℓ' + WildCatInstance.toWildCat WildCatWS x = x + WildCatInstance.mbIsWildGroupoid WildCatWS = nothing private - module WC-WS = WildStr WildCatWS + module WC-WS = WildCatInstance WildCatWS macro solveWildCat : R.Term → R.Term → R.TC Unit solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) + +module Cat-Solver ℓ ℓ' where + + Cat→WildCat : Category ℓ ℓ' → WildCat ℓ ℓ' + WildCat.ob (Cat→WildCat x) = Category.ob x + WildCat.Hom[_,_] (Cat→WildCat x) = Category.Hom[_,_] x + WildCat.id (Cat→WildCat x) = Category.id x + WildCat._⋆_ (Cat→WildCat x) = Category._⋆_ x + WildCat.⋆IdL (Cat→WildCat x) = Category.⋆IdL x + WildCat.⋆IdR (Cat→WildCat x) = Category.⋆IdR x + WildCat.⋆Assoc (Cat→WildCat x) = Category.⋆Assoc x + + CatWS : WildCatInstance ℓ ℓ' + WildCatInstance.wildStr CatWS = Category ℓ ℓ' + WildCatInstance.toWildCat CatWS = Cat→WildCat + WildCatInstance.mbIsWildGroupoid CatWS = nothing + + Functor' : Category ℓ ℓ' → Category ℓ ℓ' → Type (ℓ-max ℓ ℓ') + Functor' x x₁ = WildFunctor (Cat→WildCat x) (Cat→WildCat x₁) + + private + module C-WS = WildCatInstance CatWS + + macro + solveCat : R.Term → R.Term → R.TC Unit + solveCat = C-WS.solveW (R.def (quote CatWS) ( R.unknown v∷ R.unknown v∷ [])) diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index e64929a3a6..35ff084f59 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -133,7 +133,7 @@ mbFunctorApp : R.Term → Maybe (R.Term × R.Term) mbFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t mbFunctorApp _ = nothing -record WildStr ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where +record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where no-eta-equality field wildStr : Type (ℓ-suc (ℓ-max ℓ ℓ')) From 51be7c8a252017b74b826e987bd2686111bf43d5 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Sun, 24 Mar 2024 21:16:32 +0100 Subject: [PATCH 08/20] path-solver --- Cubical/Data/List/Properties.agda | 22 ++ Cubical/Foundations/GroupoidLaws.agda | 7 + Cubical/Foundations/Path.agda | 28 ++ Cubical/Tactics/PathSolver/Example.agda | 90 +++++ Cubical/Tactics/PathSolver/Solver.agda | 504 ++++++++++++++++++++++++ 5 files changed, 651 insertions(+) create mode 100644 Cubical/Tactics/PathSolver/Example.agda create mode 100644 Cubical/Tactics/PathSolver/Solver.agda diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index 50063cd2ab..ce4ad2fa87 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -5,9 +5,11 @@ open import Agda.Builtin.List open import Cubical.Core.Everything open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat +open import Cubical.Data.Maybe open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Relation.Nullary @@ -192,3 +194,23 @@ intersperse a (x ∷ xs) = x ∷ a ∷ intersperse a xs join : List (List A) → List A join [] = [] join (x ∷ xs) = x ++ join xs + + +rot : List A → List A +rot [] = [] +rot (x ∷ xs) = xs ∷ʳ x + +rotN : ℕ → List A → List A +rotN n = iter n rot + +module _ {A : Type ℓ} (_≟_ : Discrete A) where + + private + fa : ℕ → (xs ys : List A) → Maybe (Σ _ λ k → xs ≡ rotN k ys) + fa zero _ _ = nothing + fa (suc k) xs ys = + decRec (just ∘ ((length xs ∸ k) ,_)) + (λ _ → fa k xs ys) (discreteList _≟_ xs (rotN (length xs ∸ k) ys) ) + + findAligment : (xs ys : List A) → Maybe (Σ _ λ k → xs ≡ rotN k ys) + findAligment xs ys = fa (suc (length xs)) xs ys diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 9219752529..c632117408 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -78,6 +78,13 @@ rCancel-filler' {x = x} {y} p i j k = rCancel' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ p ⁻¹ ≡ refl rCancel' p j k = rCancel-filler' p i0 j k +midCancel : (p : y ≡ x) (q : z ≡ x) (r : w ≡ x) + → (p ∙∙ refl ∙∙ sym q) ∙ (q ∙∙ refl ∙∙ sym r) ≡ (p ∙∙ refl ∙∙ sym r) +midCancel p q r j = + (λ i → p (i ∧ j)) + ∙∙ doubleCompPath-filler p refl (sym q) (~ j) + ∙∙ λ i → hcomp (λ k → λ {(i = i0) → q (~ k ∨ j) ; (i = i1) → r (~ k) ; (j = i1) → r (~ i ∨ ~ k) }) (r i1) + lCancel : (p : x ≡ y) → p ⁻¹ ∙ p ≡ refl lCancel p = rCancel (p ⁻¹) diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index f8a3270618..3ca12efb71 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -437,3 +437,31 @@ Square→compPathΩ² {a = a} sq k i j = ; (j = i1) → a ; (k = i1) → cong (λ x → rUnit x r) (flipSquare sq) i j}) (sq j i) + +module 2-cylinder-from-square + {a₀₀ a₀₁ a₁₀ a₁₁ a₀₀' a₀₁' a₁₀' a₁₁' : A } + {a₀₋ : a₀₀ ≡ a₀₁ } {a₁₋ : a₁₀ ≡ a₁₁ } {a₋₀ : a₀₀ ≡ a₁₀ } {a₋₁ : a₀₁ ≡ a₁₁ } + {a₀₋' : a₀₀' ≡ a₀₁'} {a₁₋' : a₁₀' ≡ a₁₁'} {a₋₀' : a₀₀' ≡ a₁₀'} {a₋₁' : a₀₁' ≡ a₁₁'} + (aa'₀₀ : a₀₀ ≡ a₀₀') + where + + MissingSquare = (a₁₋ ⁻¹ ∙∙ a₋₀ ⁻¹ ∙∙ aa'₀₀ ∙∙ a₋₀' ∙∙ a₁₋') + ≡ (a₋₁ ⁻¹ ∙∙ a₀₋ ⁻¹ ∙∙ aa'₀₀ ∙∙ a₀₋' ∙∙ a₋₁') + + cyl : MissingSquare → ∀ i j → I → Partial (i ∨ ~ i ∨ j ∨ ~ j) A + + cyl c i j k = λ where + (i = i0) → doubleCompPath-filler (sym a₀₋) aa'₀₀ a₀₋' j k + (i = i1) → doubleCompPath-filler (sym a₁₋) (sym a₋₀ ∙∙ aa'₀₀ ∙∙ a₋₀') a₁₋' j k + (j = i0) → doubleCompPath-filler (sym a₋₀) aa'₀₀ a₋₀' i k + (j = i1) → compPathR→PathP∙∙ c (~ i) k + + module _ (s : MissingSquare) where + IsoSqSq' : Iso (Square a₀₋ a₁₋ a₋₀ a₋₁) (Square a₀₋' a₁₋' a₋₀' a₋₁') + Iso.fun IsoSqSq' x i j = hcomp (cyl s i j) (x i j) + Iso.inv IsoSqSq' x i j = hcomp (λ k → cyl s i j (~ k)) (x i j) + Iso.rightInv IsoSqSq' x l i j = hcomp-equivFiller (λ k → cyl s i j (~ k)) (inS (x i j)) (~ l) + Iso.leftInv IsoSqSq' x l i j = hcomp-equivFiller (cyl s i j) (inS (x i j)) (~ l) + + Sq≃Sq' : (Square a₀₋ a₁₋ a₋₀ a₋₁) ≃ (Square a₀₋' a₁₋' a₋₀' a₋₁') + Sq≃Sq' = isoToEquiv IsoSqSq' diff --git a/Cubical/Tactics/PathSolver/Example.agda b/Cubical/Tactics/PathSolver/Example.agda new file mode 100644 index 0000000000..1fe3d8dcb8 --- /dev/null +++ b/Cubical/Tactics/PathSolver/Example.agda @@ -0,0 +1,90 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Example where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure + +open import Cubical.Tactics.PathSolver.Solver + +private + variable + ℓ ℓ' : Level + +module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where + + pA pB pC : x ≡ w + pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + + pA=pB : pA ≡ pB + pA=pB = solveGroupoidDebug + + pB=pC : pB ≡ pC + pB=pC = solveGroupoidDebug + + pA=pC : pA ≡ pC + pA=pC = solveGroupoidDebug + + pA=pB∙pB=pC-≡-pA=pC : pA=pB ∙ pB=pC ≡ pA=pC + pA=pB∙pB=pC-≡-pA=pC = midCancel _ _ _ + + + sqTest : Square p (sym r ∙ refl) (p ∙ q) (q ∙ sym r) + sqTest = solveSquareDebug + +module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where + + open import Cubical.Foundations.Equiv + + SqCompEquiv : (Square p p q q) ≃ (p ∙ q ≡ q ∙ p) + SqCompEquiv = 2-cylinder-from-square.Sq≃Sq' refl solveSquareDebug + + + +open import Cubical.Algebra.Group + +-- module EM₁-Example (G : Group ℓ) (a b c : fst G) where +-- open GroupStr (snd G) + +-- data EM₁ : Type ℓ where +-- embase : EM₁ +-- emloop : ⟨ G ⟩ → embase ≡ embase +-- emcomp : (g h : ⟨ G ⟩ ) → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) + +-- double-emcomp : Square {A = EM₁} +-- (emloop b) (emloop ((a · b) · c)) +-- (sym (emloop a)) (emloop c) + +-- double-emcomp = zzz + +-- where +-- zz : (λ i → +-- hcomp +-- (doubleComp-faces (λ i₁ → emloop ((a · b) · c) (~ i₁)) +-- (emloop ((a · b) · c)) i) +-- (hcomp +-- (doubleComp-faces +-- (λ i₁ → +-- hcomp (doubleComp-faces (λ _ → embase) (λ i₂ → embase) (~ i₁)) +-- embase) +-- (λ i₁ → emloop a (~ i₁)) i) +-- (emloop a i))) +-- ≡ +-- (λ i → +-- hcomp +-- (doubleComp-faces +-- (λ i₁ → +-- hcomp (doubleComp-faces (λ _ → embase) (λ i₂ → emloop c i₂) (~ i₁)) +-- (emloop b (~ i₁))) +-- (λ i₁ → emloop c i₁) i) +-- (hcomp (doubleComp-faces (λ i₁ → emloop a (~ i₁)) (emloop b) i) +-- (emloop a i))) +-- zz = solveGroupoidDebug + +-- zzz = fst (2-cylinder-from-square.Sq≃Sq' +-- (emloop a) +-- zz) (emcomp a b ∙v emcomp (a · b) c) diff --git a/Cubical/Tactics/PathSolver/Solver.agda b/Cubical/Tactics/PathSolver/Solver.agda new file mode 100644 index 0000000000..324669b52c --- /dev/null +++ b/Cubical/Tactics/PathSolver/Solver.agda @@ -0,0 +1,504 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Solver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + +private + variable + ℓ ℓ' : Level + + + +module _ {A : Type ℓ} where + data PathCase : {a₀ a₁ : A} → a₀ ≡ a₁ → Type ℓ where + reflCase : ∀ {x} → PathCase (refl {x = x}) + compCase : ∀ {x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase (p ∙∙ q ∙∙ r) + + +module PT {A : Type ℓ} (_∼_ : A → A → Type ℓ') (∼refl : ∀ {x} → x ∼ x) + (_∼∙_ : ∀ {x y z} → x ∼ y → y ∼ z → x ∼ z) + (_∼∙∙_∼∙∙_ : ∀ {x y z w} → x ∼ y → y ∼ z → z ∼ w → x ∼ w) + (∼doubleCompPath-elim : ∀ {x y z w} → + (p : x ∼ y) → (q : y ∼ z) → (r : z ∼ w) → (p ∼∙∙ q ∼∙∙ r) ≡ (p ∼∙ q) ∼∙ r) + (∼assoc : ∀ {x y z w} → (p : x ∼ y) (q : y ∼ z) (r : z ∼ w) → p ∼∙ (q ∼∙ r) ≡ (p ∼∙ q) ∼∙ r) + where + + + + data IsPathTrm : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + isReflTrm : ∀ {x} → IsPathTrm (∼refl {x = x}) + isAtomTrm : ∀ {x y} → (p : x ∼ y) → IsPathTrm p + isCompTrm : ∀ {x y z w : _} → {p : x ∼ y} {q : y ∼ z} {r : z ∼ w} + → IsPathTrm p → IsPathTrm q → IsPathTrm r → IsPathTrm (p ∼∙∙ q ∼∙∙ r) + + + + + data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) + isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p + isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} + → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) + + data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) + consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) + + data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where + nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x + consTrmInvol : ∀ {x y z : _} → + IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z + involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → + IsPathTrmInvol x y + + + module show (showA : {a₀ a₁ : A} → (p : a₀ ∼ a₁) → String) where + showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String + showIPT isReflTrm = "refl" + showIPT (isAtomTrm x) = showA x + showIPT (isCompTrm isReflTrm x₁ x₂) = + "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" + showIPT (isCompTrm x x₁ isReflTrm) = + "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" + showIPT (isCompTrm x x₁ x₂) = + "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" + + showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String + + showIPTD nilTrmDeasRefl = "refl" + showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q + + showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String + showIPTI nilTrmInvolRefl = "refl" + showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q + showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" + + + depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} + → IsPathTrmDeas p → ℕ + depthIsPathTrmDeas nilTrmDeasRefl = zero + depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) + + hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool + hasRedexes nilTrmInvolRefl = false + hasRedexes (consTrmInvol x q) = hasRedexes x + hasRedexes (involTrmInvol x q) = true + + Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ + Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl + Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q + + IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → + IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q + IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' + IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' + IsPathTrmDeas∙ p' (consTrmDeas q' q'') = + let (_ , u) = IsPathTrmDeas∙ p' q' + in _ , consTrmDeas u q'' + + + Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} + + Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl + Invol→Deas↓ (consTrmInvol x q) = + IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) + Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x + + ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) + ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl + ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl + ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = + let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r + in _ , isCompTrmReg (isCompTrmReg pR qR) rR , + λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i + + + ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) + ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl + ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p + ⟦ isCompTrmReg p' q' ⟧da = + let (_ , qD) = ⟦ q' ⟧da + (_ , pD) = ⟦ p' ⟧da + (_ , p∙qD) = IsPathTrmDeas∙ pD qD + in _ , p∙qD + + ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) + ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da + module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) + (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where + + IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → + (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → + (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) + IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ + IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ + IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = + cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ + sym (∼assoc _ _ _) + + ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → + fst ⟦ p' ⟧da ≡ p + ⟦ isReflTrmReg ⟧da≡ = refl + ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) + ⟦ isCompTrmReg p' q' ⟧da≡ = + IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ + + daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da + daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) + + + +module _ {A : Type ℓ} where + + open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public + open PT≡ rUnit lUnit public + + + + ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ + ⟦ nilTrmInvolRefl , _ ⟧ti = refl + ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q + ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q + ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti + ⟦ involTrmInvol x q , seg j ⟧ti i = + hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i + ; (i = i1) → q (~ k ∧ ~ j) + ; (i = i0) → ⟦ x , seg j ⟧ti i0 + }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) + + ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti + ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti + + + +module _ (A : Type ℓ) (a : A) where + module PTG = PT {A = Unit} (λ _ _ → A) + (a) + (λ _ _ → a) + (λ _ _ _ → a) + (λ _ _ _ → refl) + (λ _ _ _ → refl) + +module PTrm = PTG R.Term R.unknown + +module Pℕ = PTG (Bool × ℕ) (true , 0) + +module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) + + +module _ (f : (Bool × ℕ) → R.Term) where + mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ + mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl + mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) + mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) + + +IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) +IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ + +ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ + +consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ +consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x +consInvolℕ q'@(PT.consTrmDeas x q) x₁ = + decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) + (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) + + + +ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl +ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' + +module tryPathE where + + try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) + + + tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + R.returnTC (_ , PTrm.isReflTrm) + + tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (_ , t1') ← try≡ k t1 + (_ , t2') ← try≡ k t2 + (_ , t3') ← try≡ k t3 + R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) + atom x = R.returnTC (_ , PTrm.isAtomTrm x) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) (atom t)) + +lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y +lamWrap' p i = p i + + +lamWrap : R.Term → R.Term +lamWrap t = R.def (quote lamWrap') (t v∷ []) + +unLam : R.Term → R.TC R.Term +unLam (R.lam _ (R.abs _ t)) = R.returnTC t +unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) + +appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) +appendIfUniqe x [] = R.returnTC [ x ] +appendIfUniqe x xs@(x₁ ∷ xs') = do + x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= + λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) + x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam + sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam + R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ + ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) + ( + R.catchTC + (R.extendContext "i" (varg (R.def (quote I) [])) $ + ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) + (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) + ) + +comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) +comparePathTerms x x₁ = do + x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= + λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) + x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= + λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) + sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam + R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ + ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) + ( + R.catchTC + (R.extendContext "i" (varg (R.def (quote I) [])) $ + ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) + (R.returnTC nothing) + ) + +concatUniq : List R.Term → List R.Term → R.TC (List R.Term) +concatUniq [] = R.returnTC +concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ + +indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) +indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) +indexOfAlways t xs@(x ∷ xs') = + comparePathTerms t x >>= + Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) + (λ b → R.returnTC (xs , (b , 0))) + +unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → + (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) +unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) +unMapAtoms l (PT.isAtomTrm x) = + do (l' , y) ← indexOfAlways x l + R.returnTC (l' , _ , Pℕ.isAtomTrm y) +unMapAtoms l (PT.isCompTrm e e₁ e₂) = + do (l' , _ , e') ← unMapAtoms l e + (l'' , _ , e₁') ← unMapAtoms l' e₁ + (l''' , _ , e₂') ← unMapAtoms l'' e₂ + (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) + + +unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term +unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] +unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) +unquotePathTrm (PT.isCompTrm x x₁ x₂) = + let x' = unquotePathTrm x + x₁' = unquotePathTrm x₁ + x₂' = unquotePathTrm x₂ + in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) + +module _ (l : List R.Term) where + lk : (Bool × ℕ) → R.Term + lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) + where + z = lookupWithDefault R.unknown l n + + + + mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) + mkTrmsInvol' zero _ = [] + mkTrmsInvol' (suc k) u = + let pi = ℕDeas→ℕInvol u + in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] + + mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) + mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x + + unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term + -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) + -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) + unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] + unquoteTrmInvol (PT.consTrmInvol x q) = + R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) + unquoteTrmInvol (PT.involTrmInvol x q) = + R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) + + mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) + mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) + ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x + + foldPathCompTerm : List R.Term → R.Term + foldPathCompTerm [] = R.def (quote refl) [] + foldPathCompTerm (x ∷ []) = x + foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) + + mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) + mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) + + +groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag hole = do + + (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + (r0 , r0') ← tryPathE.try≡ 100 t0 + (r1 , r1') ← tryPathE.try≡ 100 t1 + + + (aL' , (_ , e0)) ← unMapAtoms [] r0' + (aL , (_ , e1)) ← unMapAtoms aL' r1' + let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + + let dbgInfo = (R.strErr "LHS: ") ∷ (R.termErr $ t0) + ∷ (R.strErr "\n") + ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + ∷ (R.strErr "\n") + ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + ∷ (R.strErr "\n") + ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + ∷ (R.strErr "\n") + ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + ∷ (R.strErr "\n") + ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + ∷ (R.strErr "\n") + ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + ∷ (R.strErr "\n") + ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) + + ∷ (R.strErr "\n") + ∷ (niceAtomList aL) + + + let (_ , iP0) = mkTrmInvol aL e0deas + (eqs1 , iP1) = mkTrmInvol aL e1deas + + let dbgInfo = dbgInfo ++ ((R.strErr "\n") ∷ niceEqsList eqs1) + + + let pa0 = R.def (quote _∙_) + (R.def (quote daSingl) ((unquotePathTrm r0') v∷ []) + v∷ iP0 v∷ [] ) + pa1 = R.def (quote _∙_) + (R.def (quote daSingl) ((unquotePathTrm r1') v∷ []) + v∷ iP1 v∷ [] ) + let final = (R.def (quote _∙∙_∙∙_) + (pa0 + v∷ R.def (quote refl) [] + v∷ R.def (quote sym) (pa1 v∷ []) v∷ [] )) + + + R.returnTC (final , dbgInfo) + + where + + niceEq : ℕ → Pℕ.IsPathTrmInvol _ _ → List R.ErrorPart + niceEq k x = R.strErr (primShowNat k <> " : ") + ∷ R.strErr (PℕS.showIPTI x) + ∷ [ R.strErr "\n" ] + + niceEqsList' : ℕ → List (Pℕ.IsPathTrmInvol _ _) → List R.ErrorPart + niceEqsList' k [] = [] + niceEqsList' k (x ∷ xs) = + niceEq k x ++ niceEqsList' (suc k) xs + + niceEqsList = niceEqsList' 0 + +groupoidSolverMain : Bool → R.Term → R.TC Unit +groupoidSolverMain debugFlag hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + +squareSolverMain : Bool → R.Term → R.TC Unit +squareSolverMain debugFlag hole = do + ty ← R.inferType hole >>= wait-for-type + hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta + + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole' solution) + (R.typeError msg) + + R.catchTC + (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) + (R.typeError [ R.strErr "xxx" ] ) + where + extractMeta : R.Term → R.TC R.Term + extractMeta (R.def _ tl) = matchVarg tl + extractMeta t = + R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) + +macro + solveGroupoidDebug : R.Term → R.TC Unit + solveGroupoidDebug = groupoidSolverMain true + + solveSquareDebug : R.Term → R.TC Unit + solveSquareDebug = squareSolverMain false + + solveGroupoid : R.Term → R.TC Unit + solveGroupoid = groupoidSolverMain true + + solveSquare : R.Term → R.TC Unit + solveSquare = squareSolverMain false From 2db814accc2318ac5b8fe99cf033ec099cecd0b3 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 25 Mar 2024 21:22:37 +0100 Subject: [PATCH 09/20] wip on paths --- Cubical/Reflection/Base.agda | 4 + Cubical/Tactics/PathSolver/Example.agda | 55 +- Cubical/Tactics/PathSolver/Solver2.agda | 647 ++++++++++++++++++ Cubical/Tactics/Reflection.agda | 4 + Cubical/Tactics/WildCatSolver/Reflection.agda | 2 + 5 files changed, 671 insertions(+), 41 deletions(-) create mode 100644 Cubical/Tactics/PathSolver/Solver2.agda diff --git a/Cubical/Reflection/Base.agda b/Cubical/Reflection/Base.agda index dd81312c08..9bea9626c3 100644 --- a/Cubical/Reflection/Base.agda +++ b/Cubical/Reflection/Base.agda @@ -7,6 +7,7 @@ Some basic utilities for reflection module Cubical.Reflection.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Data.List.Base open import Cubical.Data.Nat.Base @@ -19,6 +20,9 @@ _<|>_ = R.catchTC _>>_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → R.TC A → R.TC B → R.TC B f >> g = f >>= λ _ → g +pure : ∀ {ℓ} {A : Type ℓ} → A → R.TC A +pure = R.returnTC + infixl 4 _>>=_ _>>_ _<|>_ liftTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → R.TC A → R.TC B diff --git a/Cubical/Tactics/PathSolver/Example.agda b/Cubical/Tactics/PathSolver/Example.agda index 1fe3d8dcb8..99c852fba9 100644 --- a/Cubical/Tactics/PathSolver/Example.agda +++ b/Cubical/Tactics/PathSolver/Example.agda @@ -47,44 +47,17 @@ module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where open import Cubical.Algebra.Group --- module EM₁-Example (G : Group ℓ) (a b c : fst G) where --- open GroupStr (snd G) - --- data EM₁ : Type ℓ where --- embase : EM₁ --- emloop : ⟨ G ⟩ → embase ≡ embase --- emcomp : (g h : ⟨ G ⟩ ) → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) - --- double-emcomp : Square {A = EM₁} --- (emloop b) (emloop ((a · b) · c)) --- (sym (emloop a)) (emloop c) - --- double-emcomp = zzz - --- where --- zz : (λ i → --- hcomp --- (doubleComp-faces (λ i₁ → emloop ((a · b) · c) (~ i₁)) --- (emloop ((a · b) · c)) i) --- (hcomp --- (doubleComp-faces --- (λ i₁ → --- hcomp (doubleComp-faces (λ _ → embase) (λ i₂ → embase) (~ i₁)) --- embase) --- (λ i₁ → emloop a (~ i₁)) i) --- (emloop a i))) --- ≡ --- (λ i → --- hcomp --- (doubleComp-faces --- (λ i₁ → --- hcomp (doubleComp-faces (λ _ → embase) (λ i₂ → emloop c i₂) (~ i₁)) --- (emloop b (~ i₁))) --- (λ i₁ → emloop c i₁) i) --- (hcomp (doubleComp-faces (λ i₁ → emloop a (~ i₁)) (emloop b) i) --- (emloop a i))) --- zz = solveGroupoidDebug - --- zzz = fst (2-cylinder-from-square.Sq≃Sq' --- (emloop a) --- zz) (emcomp a b ∙v emcomp (a · b) c) +module EM₁-Example (G : Group ℓ) (a b c : fst G) where + open GroupStr (snd G) + + data EM₁ : Type ℓ where + embase : EM₁ + emloop : ⟨ G ⟩ → embase ≡ embase + emcomp : (g h : ⟨ G ⟩ ) → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) + + double-emcomp : Square {A = EM₁} + (emloop b) (emloop ((a · b) · c)) + (sym (emloop a)) (emloop c) + + double-emcomp = fst (2-cylinder-from-square.Sq≃Sq' + (emloop a) solveSquareDebug) (emcomp a b ∙v emcomp (a · b) c) diff --git a/Cubical/Tactics/PathSolver/Solver2.agda b/Cubical/Tactics/PathSolver/Solver2.agda new file mode 100644 index 0000000000..fc8e7f0abf --- /dev/null +++ b/Cubical/Tactics/PathSolver/Solver2.agda @@ -0,0 +1,647 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Solver2 where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + +private + variable + ℓ ℓ' : Level + + +infixr 5 _∷Tω_ + +data [Typeₙ] : Typeω where + [] : [Typeₙ] + _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] + +ℓ[Typeₙ] : [Typeₙ] → Level +ℓ[Typeₙ] [] = ℓ-zero +ℓ[Typeₙ] (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (ℓ[Typeₙ] x₁) + +lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (ℓ[Typeₙ] Ts) +lookupTₙ [] x = Unit +lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {ℓ[Typeₙ] Ts} x₁ +lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) + +data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) + compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) + congCase : ∀ {ℓ ℓ' A A'} (f : A → A') {x y} + → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) + + +-- module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} +-- (p : Path {ℓ} A x y) where + +-- pathCaseCongTest : PathCase (cong f p) +-- pathCaseCongTest = congCase f λ z → p z + +data PathExprω : {ℓ : Level} (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflExpr : ∀ {ℓ A x} → PathExprω {ℓ} A (refl {x = x}) + atomExpr : ∀ {ℓ A x y} → (p : x ≡ y) → PathExprω {ℓ} A p + compExpr : ∀ {ℓ A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} + → PathExprω {ℓ} A p → PathExprω {ℓ} A q → PathExprω {ℓ} A r + → PathExprω {ℓ} A (p ∙∙ q ∙∙ r) + congExpr : ∀ {ℓ ℓ' A A' x y} → (f : A → A') {p : x ≡ y} + → PathExprω {ℓ} A p + → PathExprω {ℓ'} A' (cong f p) + + +data PathExpr {ℓ : Level} : (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Type (ℓ-suc ℓ) where + reflExpr : ∀ {A x} → PathExpr A (refl {x = x}) + atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A p + compExpr : ∀ {A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} + → PathExpr A p → PathExpr A q → PathExpr A r + → PathExpr A (p ∙∙ q ∙∙ r) + congExpr : ∀ {A A' x y} → (f : A → A') {p : x ≡ y} + → PathExpr {ℓ} A {x} {y} p + → PathExpr A' (cong f p) + + + +PEω⟦_⟧ : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → + PathExprω A p → a₀ ≡ a₁ +PEω⟦ reflExpr ⟧ = refl +PEω⟦ atomExpr p ⟧ = p +PEω⟦ compExpr x x₁ x₂ ⟧ = PEω⟦ x ⟧ ∙∙ PEω⟦ x₁ ⟧ ∙∙ PEω⟦ x₂ ⟧ +PEω⟦ congExpr f x ⟧ = cong f PEω⟦ x ⟧ + + +PE⟦_⟧ : {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → + PathExpr A p → a₀ ≡ a₁ +PE⟦ reflExpr ⟧ = refl +PE⟦ atomExpr p ⟧ = p +PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ +PE⟦ congExpr f x ⟧ = cong f PE⟦ x ⟧ + + +record ΣPathExprω : Typeω where + field + ℓa : Level + A : Type ℓa + a₀ a₁ : A + p : a₀ ≡ a₁ + expr : PathExprω A p + + ΣPEω⟦_⟧ : a₀ ≡ a₁ + ΣPEω⟦_⟧ = PEω⟦ expr ⟧ + +record ΣPathExpr {ℓ} : Type (ℓ-suc ℓ) where + constructor Σpe + field + {A} : Type ℓ + {a₀ a₁} : A + {p} : a₀ ≡ a₁ + expr : PathExpr A p + + ΣPE⟦_⟧ : a₀ ≡ a₁ + ΣPE⟦_⟧ = PE⟦ expr ⟧ + + + +module PathTrm (A B : Type ℓ) where + data PathTrm : Type ℓ where + reflTrm : PathTrm + atomTrm : A → PathTrm + compTrm : PathTrm → PathTrm → PathTrm → PathTrm + congTrm : B → PathTrm → PathTrm + + module showPathTrm (showA : A → String) (showB : B → String) where + showPT : PathTrm → String + showPT reflTrm = "refl" + showPT (atomTrm x) = showA x + showPT (compTrm x x₁ x₂) = + "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" + showPT (congTrm x x₁) = + "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" + + +module _ {ℓ ℓ'} + {A B : Type ℓ} + {A' B' : Type ℓ'} + (f : A → R.TC A') + (g : B → R.TC B') where + open PathTrm + mmapPT : PathTrm A B → R.TC $ PathTrm A' B' + mmapPT reflTrm = pure reflTrm + mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ + mmapPT (compTrm x x₁ x₂) = + ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ + mmapPT (congTrm x x₁) = + ⦇ congTrm (g x) (mmapPT x₁) ⦈ + + +module RTrm = PathTrm R.Term R.Term +module StrTrm = PathTrm String String + +-- unqouteRTrm : ∀ {ℓ} → RTrm.PathTrm → R.TC (ΣPathExpr {ℓ}) +-- unqouteRTrm PathTrm.reflTrm = pure $ Σpe {A = Unit*} reflExpr +-- unqouteRTrm (PathTrm.atomTrm x) = +-- ⦇ (Σpe ∘ atomExpr) {!x!} ⦈ +-- unqouteRTrm (PathTrm.compTrm x x₁ x₂) = {!!} +-- unqouteRTrm (PathTrm.congTrm x x₁) = {!!} + +showRTrmTC : RTrm.PathTrm → R.TC String +showRTrmTC = + mmapPT + (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +module tryPathE where + + try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm) + + + tryRefl : R.Term → R.TC (RTrm.PathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + R.returnTC (RTrm.reflTrm) + + tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (t1') ← try≡ k t1 + (t2') ← try≡ k t2 + (t3') ← try≡ k t3 + R.returnTC (RTrm.compTrm t1' t2' t3') + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + atom : R.Term → R.TC (RTrm.PathTrm) + atom x = R.returnTC (RTrm.atomTrm x) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) (atom t)) + + + + + + + + + + -- data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + -- isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) + -- isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p + -- isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} + -- → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) + + -- data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + -- nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) + -- consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) + + -- data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where + -- nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x + -- consTrmInvol : ∀ {x y z : _} → + -- IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z + -- involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → + -- IsPathTrmInvol x y + + +-- module show (showA : ∀ {ℓ} A {a₀ a₁ : A} → (p : a₀ ≡ a₁) → String) where +-- showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String +-- showIPT isReflTrm = "refl" +-- showIPT (isAtomTrm x) = showA x +-- showIPT (isCompTrm isReflTrm x₁ x₂) = +-- "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" +-- showIPT (isCompTrm x x₁ isReflTrm) = +-- "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" +-- showIPT (isCompTrm x x₁ x₂) = +-- "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" + +-- -- showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String + +-- -- showIPTD nilTrmDeasRefl = "refl" +-- -- showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q + +-- -- showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String +-- -- showIPTI nilTrmInvolRefl = "refl" +-- -- showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q +-- -- showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" + + +-- -- depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} +-- -- → IsPathTrmDeas p → ℕ +-- -- depthIsPathTrmDeas nilTrmDeasRefl = zero +-- -- depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) + +-- -- hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool +-- -- hasRedexes nilTrmInvolRefl = false +-- -- hasRedexes (consTrmInvol x q) = hasRedexes x +-- -- hasRedexes (involTrmInvol x q) = true + +-- -- Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ +-- -- Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl +-- -- Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q + +-- -- IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → +-- -- IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q +-- -- IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' +-- -- IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' +-- -- IsPathTrmDeas∙ p' (consTrmDeas q' q'') = +-- -- let (_ , u) = IsPathTrmDeas∙ p' q' +-- -- in _ , consTrmDeas u q'' + + +-- -- Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} + +-- -- Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl +-- -- Invol→Deas↓ (consTrmInvol x q) = +-- -- IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) +-- -- Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x + +-- -- ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) +-- -- ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl +-- -- ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl +-- -- ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = +-- -- let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r +-- -- in _ , isCompTrmReg (isCompTrmReg pR qR) rR , +-- -- λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i + + +-- -- ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) +-- -- ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl +-- -- ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p +-- -- ⟦ isCompTrmReg p' q' ⟧da = +-- -- let (_ , qD) = ⟦ q' ⟧da +-- -- (_ , pD) = ⟦ p' ⟧da +-- -- (_ , p∙qD) = IsPathTrmDeas∙ pD qD +-- -- in _ , p∙qD + +-- -- ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) +-- -- ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da +-- -- module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) +-- -- (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where + +-- -- IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → +-- -- (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → +-- -- (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) +-- -- IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ +-- -- IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ +-- -- IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = +-- -- cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ +-- -- sym (∼assoc _ _ _) + +-- -- ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → +-- -- fst ⟦ p' ⟧da ≡ p +-- -- ⟦ isReflTrmReg ⟧da≡ = refl +-- -- ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) +-- -- ⟦ isCompTrmReg p' q' ⟧da≡ = +-- -- IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ + +-- -- daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da +-- -- daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) + + + +-- -- module _ {A : Type ℓ} where + +-- -- open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public +-- -- open PT≡ rUnit lUnit public + + + +-- -- ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ +-- -- ⟦ nilTrmInvolRefl , _ ⟧ti = refl +-- -- ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q +-- -- ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q +-- -- ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti +-- -- ⟦ involTrmInvol x q , seg j ⟧ti i = +-- -- hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i +-- -- ; (i = i1) → q (~ k ∧ ~ j) +-- -- ; (i = i0) → ⟦ x , seg j ⟧ti i0 +-- -- }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) + +-- -- ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti +-- -- ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti + + + +-- module _ (A : Type ℓ) (a : A) where +-- module PTG = PT {A = Unit} (λ _ _ → A) +-- (a) +-- (λ _ _ → a) +-- (λ _ _ _ → a) +-- (λ _ _ _ → refl) +-- (λ _ _ _ → refl) + +-- module PTrm = PTG R.Term R.unknown + +-- module Pℕ = PTG (Bool × ℕ) (true , 0) + +-- module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) + + +-- -- module _ (f : (Bool × ℕ) → R.Term) where +-- -- mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ +-- -- mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl +-- -- mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) +-- -- mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) + + +-- IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) +-- IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ + +-- -- ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ + +-- -- consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ +-- -- consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x +-- -- consInvolℕ q'@(PT.consTrmDeas x q) x₁ = +-- -- decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) +-- -- (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) + + + +-- -- ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl +-- -- ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' + +-- module tryPathE where + +-- try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) + + +-- tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- tryRefl t = do +-- _ ← R.noConstraints $ R.checkType +-- (R.con (quote reflCase) []) +-- (R.def (quote PathCase) ([ varg t ])) +-- R.returnTC (_ , PTrm.isReflTrm) + +-- tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] +-- tryComp (suc k) t = do +-- tm ← R.noConstraints $ R.checkType +-- (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) +-- (R.def (quote PathCase) ([ varg t ])) +-- (t1 , t2 , t3) ← h tm +-- (_ , t1') ← try≡ k t1 +-- (_ , t2') ← try≡ k t2 +-- (_ , t3') ← try≡ k t3 +-- R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') + +-- where + +-- h : R.Term → R.TC (R.Term × R.Term × R.Term) +-- h (R.con (quote compCase) l) = match3Vargs l +-- h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + +-- atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- atom x = R.returnTC (_ , PTrm.isAtomTrm x) + + +-- try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] +-- try≡ (suc k) t = +-- R.catchTC +-- (tryRefl t) +-- (R.catchTC (tryComp k t) (atom t)) + +-- lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y +-- lamWrap' p i = p i + + +-- lamWrap : R.Term → R.Term +-- lamWrap t = R.def (quote lamWrap') (t v∷ []) + +-- unLam : R.Term → R.TC R.Term +-- unLam (R.lam _ (R.abs _ t)) = R.returnTC t +-- unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) + +-- appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) +-- appendIfUniqe x [] = R.returnTC [ x ] +-- appendIfUniqe x xs@(x₁ ∷ xs') = do +-- x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) +-- x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) +-- ( +-- R.catchTC +-- (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) +-- (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) +-- ) + +-- comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) +-- comparePathTerms x x₁ = do +-- x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) +-- x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) +-- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) +-- ( +-- R.catchTC +-- (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) +-- (R.returnTC nothing) +-- ) + +-- concatUniq : List R.Term → List R.Term → R.TC (List R.Term) +-- concatUniq [] = R.returnTC +-- concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ + +-- indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) +-- indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) +-- indexOfAlways t xs@(x ∷ xs') = +-- comparePathTerms t x >>= +-- Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) +-- (λ b → R.returnTC (xs , (b , 0))) + +-- unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → +-- (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) +-- unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) +-- unMapAtoms l (PT.isAtomTrm x) = +-- do (l' , y) ← indexOfAlways x l +-- R.returnTC (l' , _ , Pℕ.isAtomTrm y) +-- unMapAtoms l (PT.isCompTrm e e₁ e₂) = +-- do (l' , _ , e') ← unMapAtoms l e +-- (l'' , _ , e₁') ← unMapAtoms l' e₁ +-- (l''' , _ , e₂') ← unMapAtoms l'' e₂ +-- (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) + + +-- -- unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term +-- -- unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] +-- -- unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) +-- -- unquotePathTrm (PT.isCompTrm x x₁ x₂) = +-- -- let x' = unquotePathTrm x +-- -- x₁' = unquotePathTrm x₁ +-- -- x₂' = unquotePathTrm x₂ +-- -- in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) + +-- -- module _ (l : List R.Term) where +-- -- lk : (Bool × ℕ) → R.Term +-- -- lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) +-- -- where +-- -- z = lookupWithDefault R.unknown l n + + + +-- -- mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) +-- -- mkTrmsInvol' zero _ = [] +-- -- mkTrmsInvol' (suc k) u = +-- -- let pi = ℕDeas→ℕInvol u +-- -- in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] + +-- -- mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) +-- -- mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x + +-- -- unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term +-- -- -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) +-- -- -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) +-- -- unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] +-- -- unquoteTrmInvol (PT.consTrmInvol x q) = +-- -- R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) +-- -- unquoteTrmInvol (PT.involTrmInvol x q) = +-- -- R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) + +-- -- mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) +-- -- mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) +-- -- ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x + +-- -- foldPathCompTerm : List R.Term → R.Term +-- -- foldPathCompTerm [] = R.def (quote refl) [] +-- -- foldPathCompTerm (x ∷ []) = x +-- -- foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) + +-- -- mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) +-- -- mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) + + +groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag hole = do + + (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + r0 ← tryPathE.try≡ 100 t0 + r1 ← tryPathE.try≡ 100 t1 + + + -- (aL' , (_ , e0)) ← unMapAtoms [] r0' + -- (aL , (_ , e1)) ← unMapAtoms aL' r1' + -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + show0 ← showRTrmTC r0 + show1 ← showRTrmTC r1 + + let dbgInfo = ("LHS: ") ∷ₑ show0 ∷nl + ("RHS: ") ∷ₑ show1 ∷nl [] + -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) + + -- ∷ (R.strErr "\n") + -- (niceAtomList aL) + + R.typeError dbgInfo + + +groupoidSolverMain : Bool → R.Term → R.TC Unit +groupoidSolverMain debugFlag hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + +squareSolverMain : Bool → R.Term → R.TC Unit +squareSolverMain debugFlag hole = do + ty ← R.inferType hole >>= wait-for-type + hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta + + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole' solution) + (R.typeError msg) + + R.catchTC + (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) + (R.typeError [ R.strErr "xxx" ] ) + where + extractMeta : R.Term → R.TC R.Term + extractMeta (R.def _ tl) = matchVarg tl + extractMeta t = + R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) + +macro + solveGroupoidDebug : R.Term → R.TC Unit + solveGroupoidDebug = groupoidSolverMain true + + -- solveSquareDebug : R.Term → R.TC Unit + -- solveSquareDebug = squareSolverMain false + + -- solveGroupoid : R.Term → R.TC Unit + -- solveGroupoid = groupoidSolverMain true + + -- solveSquare : R.Term → R.TC Unit + -- solveSquare = squareSolverMain false + + +module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where + + pA pB pC : x ≡ w + pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + + pA=pB : pA ≡ pB + pA=pB = solveGroupoidDebug diff --git a/Cubical/Tactics/Reflection.agda b/Cubical/Tactics/Reflection.agda index ff96225e8c..802ffd21e9 100644 --- a/Cubical/Tactics/Reflection.agda +++ b/Cubical/Tactics/Reflection.agda @@ -29,6 +29,10 @@ f <$> t = t >>= λ x → returnTC (f x) _<*>_ : ∀ {ℓ ℓ'} {A : Type ℓ}{B : Type ℓ'} → TC (A → B) → TC A → TC B s <*> t = s >>= λ f → t >>= λ x → returnTC (f x) +_>=>_ : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} → + (A → TC B) → (B → TC C) → A → TC C +(x >=> x₁) x₂ = x x₂ >>= x₁ + wait-for-args : List (Arg Term) → TC (List (Arg Term)) wait-for-type : Term → TC Term diff --git a/Cubical/Tactics/WildCatSolver/Reflection.agda b/Cubical/Tactics/WildCatSolver/Reflection.agda index 414bdfc61b..288e3c99de 100644 --- a/Cubical/Tactics/WildCatSolver/Reflection.agda +++ b/Cubical/Tactics/WildCatSolver/Reflection.agda @@ -180,6 +180,8 @@ _∷ₑ_ ⦃ tep ⦄ x = (toErrorPart tep x) ∷_ _++ₑ_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → List A → List R.ErrorPart → List R.ErrorPart _++ₑ_ ⦃ tep ⦄ x = (map (toErrorPart tep) x) ++_ +[_]ₑ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart +[_]ₑ = _∷ₑ [] instance ToErrorPartString : ToErrorPart String From bfca4a8c626c846e331f28f4199845af4ba532cb Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Thu, 28 Mar 2024 08:43:32 +0100 Subject: [PATCH 10/20] wip --- .../EilenbergMacLane/Groups/KleinBottle.agda | 13 +- Cubical/Data/List/Properties.agda | 4 + .../Homotopy/EilenbergMacLane/Properties.agda | 1917 +++++++++-------- Cubical/Tactics/GroupSolver/Example.agda | 2 +- .../Tactics/PathSolver/ReflectionTest.agda | 50 + Cubical/Tactics/PathSolver/Solver2.agda | 14 +- Cubical/Tactics/PathSolver/Solver3.agda | 806 +++++++ Cubical/Tactics/PathSolver/Solver4.agda | 877 ++++++++ Cubical/Tactics/PathSolver/Solver5.agda | 876 ++++++++ Cubical/Tactics/Reflection.agda | 6 + Cubical/Tactics/WildCatSolver/Reflection.agda | 7 + 11 files changed, 3600 insertions(+), 972 deletions(-) create mode 100644 Cubical/Tactics/PathSolver/ReflectionTest.agda create mode 100644 Cubical/Tactics/PathSolver/Solver3.agda create mode 100644 Cubical/Tactics/PathSolver/Solver4.agda create mode 100644 Cubical/Tactics/PathSolver/Solver5.agda diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index 816f326ec8..50d46656e9 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -47,6 +47,8 @@ open import Cubical.HITs.EilenbergMacLane1 renaming (elimProp to elimPropEM1 ; e open import Cubical.HITs.Susp open import Cubical.HITs.RPn +open import Cubical.Tactics.PathSolver.Solver4 + open AbGroupStr private @@ -359,13 +361,10 @@ fst (H¹[K²,G]≅G×H¹[RP²,G] G) = isoToEquiv mainIso F→ f (square i j) = help i (~ j) where help : cong f (sym line1) ≡ cong f line1 - help = - lUnit (cong f (sym line1)) - ∙ cong (_∙ cong f (sym line1)) - (sym (lCancel _)) - ∙ sym (assoc _ _ _) - ∙ cong (sym (cong f line2) ∙_) (isCommΩEM-base 0 _ _ _) - ∙ PathP→compPathL (λ i j → f (square j i)) + help = + ≡! + ∙∙ cong (sym (cong f line2) ∙_) (isCommΩEM-base 0 _ _ _) + ∙∙ PathP→compPathL (λ i j → f (square j i)) F← : (g : fst G) → (RP² → EM G 1) → KleinBottle → EM G 1 F← g f point = f point diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index ce4ad2fa87..8b1fe4f51c 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -214,3 +214,7 @@ module _ {A : Type ℓ} (_≟_ : Discrete A) where findAligment : (xs ys : List A) → Maybe (Σ _ λ k → xs ≡ rotN k ys) findAligment xs ys = fa (suc (length xs)) xs ys + +zipWithIndex : List A → List (ℕ × A) +zipWithIndex [] = [] +zipWithIndex (x ∷ xs) = (zero , x) ∷ map (map-fst suc) (zipWithIndex xs) diff --git a/Cubical/Homotopy/EilenbergMacLane/Properties.agda b/Cubical/Homotopy/EilenbergMacLane/Properties.agda index 953ce48309..bcd227d20a 100644 --- a/Cubical/Homotopy/EilenbergMacLane/Properties.agda +++ b/Cubical/Homotopy/EilenbergMacLane/Properties.agda @@ -38,6 +38,8 @@ open import Cubical.HITs.EilenbergMacLane1 renaming (rec to EMrec ; elim to EM₁elim) open import Cubical.HITs.Susp +open import Cubical.Tactics.PathSolver.Solver4 + private variable ℓ ℓ' ℓ'' : Level @@ -75,9 +77,7 @@ module _ (Ĝ : Group ℓ) where emloop-id : emloop 1g ≡ refl emloop-id = - emloop 1g ≡⟨ rUnit (emloop 1g) ⟩ - emloop 1g ∙ refl ≡⟨ cong (emloop 1g ∙_) (rCancel (emloop 1g) ⁻¹) ⟩ - emloop 1g ∙ (emloop 1g ∙ (emloop 1g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ + emloop 1g ≡⟨ zz ⟩ (emloop 1g ∙ emloop 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (_∙ emloop 1g ⁻¹) ((emloop-comp Ĝ 1g 1g) ⁻¹) ⟩ emloop (1g · 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (λ g → emloop {Group = Ĝ} g @@ -85,7 +85,10 @@ module _ (Ĝ : Group ℓ) where (·IdR 1g) ⟩ emloop 1g ∙ (emloop 1g) ⁻¹ ≡⟨ rCancel (emloop 1g) ⟩ refl ∎ - + where + zz : emloop 1g ≡ (emloop 1g ∙ emloop 1g) ∙ emloop 1g ⁻¹ + zz = solveGroupoidDebug ((EM₁ Ĝ) ∷Tω [Tω]) [fn] + emloop-inv : (g : G) → emloop (inv g) ≡ (emloop g) ⁻¹ emloop-inv g = emloop (inv g) ≡⟨ rUnit (emloop (inv g)) ⟩ @@ -101,956 +104,956 @@ module _ (Ĝ : Group ℓ) where refl ∙ (emloop g) ⁻¹ ≡⟨ (lUnit ((emloop g) ⁻¹)) ⁻¹ ⟩ (emloop g) ⁻¹ ∎ - isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) - isGroupoidEM₁ = emsquash - - --------- Ω (EM₁ G) ≃ G --------- - - {- since we write composition in diagrammatic order, - and function composition in the other order, - we need right multiplication here -} - rightEquiv : (g : G) → G ≃ G - rightEquiv g = isoToEquiv isom - where - isom : Iso G G - isom .Iso.fun = _· g - isom .Iso.inv = _· inv g - isom .Iso.rightInv h = - (h · inv g) · g ≡⟨ (·Assoc h (inv g) g) ⁻¹ ⟩ - h · inv g · g ≡⟨ cong (h ·_) (·InvL g) ⟩ - h · 1g ≡⟨ ·IdR h ⟩ h ∎ - isom .Iso.leftInv h = - (h · g) · inv g ≡⟨ (·Assoc h g (inv g)) ⁻¹ ⟩ - h · g · inv g ≡⟨ cong (h ·_) (·InvR g) ⟩ - h · 1g ≡⟨ ·IdR h ⟩ h ∎ - - compRightEquiv : (g h : G) - → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) - compRightEquiv g h = equivEq (funExt (λ x → (·Assoc x g h) ⁻¹)) - - CodesSet : EM₁ Ĝ → hSet ℓ - CodesSet = EMrec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp - where - RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) - RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) - - lemma₁ : (g h : G) → Square - (ua (rightEquiv g)) (ua (rightEquiv (g · h))) - refl (ua (rightEquiv h)) - lemma₁ g h = invEq - (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) - refl (ua (rightEquiv h))) - (ua (rightEquiv g) ∙ ua (rightEquiv h) - ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ - ua (compEquiv (rightEquiv g) (rightEquiv h)) - ≡⟨ cong ua (compRightEquiv g h) ⟩ - ua (rightEquiv (g · h)) ∎) - - REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) - REComp g h = ΣSquareSet (λ _ → isProp→isSet isPropIsSet) (lemma₁ g h) - Codes : EM₁ Ĝ → Type ℓ - Codes x = CodesSet x .fst - - encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x - encode x p = subst Codes p 1g - - decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x - decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) - emloop λ g → ua→ λ h → emcomp h g - - decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p - decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) - (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ - emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p - - encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c - encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) - λ g → encode embase (decode embase g) ≡⟨ refl ⟩ - encode embase (emloop g) ≡⟨ refl ⟩ - transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ - 1g · g ≡⟨ ·IdL g ⟩ - g ∎ - - ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G - Iso.fun ΩEM₁Iso = encode embase - Iso.inv ΩEM₁Iso = emloop - Iso.rightInv ΩEM₁Iso = encode-decode embase - Iso.leftInv ΩEM₁Iso = decode-encode embase - - ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G - ΩEM₁≡ = isoToPath ΩEM₁Iso - ---------- Ω (EMₙ₊₁ G) ≃ EMₙ G --------- -module _ {G : AbGroup ℓ} where - open AbGroupStr (snd G) - renaming (_+_ to _+G_ ; -_ to -G_ ; +Assoc to +AssocG) - - CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓ (3 + n) - CODE n = - Trunc.elim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) - λ { north → EM G (suc n) , hLevelEM G (suc n) - ; south → EM G (suc n) , hLevelEM G (suc n) - ; (merid a i) → fib n a i} - where - fib : (n : ℕ) → (a : EM-raw G (suc n)) - → Path (TypeOfHLevel _ (3 + n)) - (EM G (suc n) , hLevelEM G (suc n)) - (EM G (suc n) , hLevelEM G (suc n)) - fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) - (isoToPath (addIso 1 a)) - fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) - (isoToPath (addIso (suc (suc n)) ∣ a ∣)) - - decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x - decode' n = - Trunc.elim (λ _ → isOfHLevelΠ (4 + n) - λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) - λ { north → f n - ; south → g n - ; (merid a i) → main a i} - where - f : (n : ℕ) → _ - f n = σ-EM' n - - g : (n : ℕ) → EM G (suc n) → ∣ ptEM-raw ∣ ≡ ∣ south ∣ - g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptEM-raw) - - lem₂ : (n : ℕ) (a x : _) - → Path (Path (EM G (suc (suc n))) _ _) - ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) - (g n (EM-raw→EM G (suc n) x)) - lem₂ zero a x = - cong (f zero x ∙_) - (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) - ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _)) - lem₂ (suc n) a x = - cong (f (suc n) ∣ x ∣ ∙_) - ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) - ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _))) - - lem : (n : ℕ) (x a : EM-raw G (suc n)) - → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) - (EM-raw→EM G (suc n) x)) - ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) - lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) - ∙∙ σ-EM'-hom zero x (-ₖ a) - ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) - lem (suc n) x a = - cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) - ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) - ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) - - main : (a : _) - → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst - → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) - main a = - toPathP (funExt - (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) - λ x → - ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (i ∨ j) ∣) - i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) - ∙∙ sym (∙assoc _ _ _) - ∙∙ lem₂ n a x)) - - encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst - encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) - - decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) - → decode' n x (encode' n x p) ≡ p - decode'-encode' zero x = - J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) - (σ-EM'-0ₖ 0) - decode'-encode' (suc n) x = - J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) - (σ-EM'-0ₖ (suc n)) - - encode'-decode' : (n : ℕ) (x : _) - → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x - encode'-decode' zero x = - cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) - ∙∙ substComposite (λ x → CODE zero x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) embase - ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) - (λ i → lUnitₖ 1 (transportRefl x i) i) - ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) - encode'-decode' (suc n) = - Trunc.elim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) - ∙∙ substComposite (λ x → CODE (suc n) x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) (0ₖ (2 + n)) - ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptEM-raw (~ i) ∣ₕ)) - (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - - Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙ G (suc n)))) - Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) - Iso.fun (Iso-EM-ΩEM+1 (suc n)) = decode' n (0ₖ (2 + n)) - Iso.inv (Iso-EM-ΩEM+1 (suc n)) = encode' n ∣ north ∣ - Iso.rightInv (Iso-EM-ΩEM+1 (suc n)) = decode'-encode' _ _ - Iso.leftInv (Iso-EM-ΩEM+1 (suc n)) = encode'-decode' _ - - EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) - EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) - - -- Some properties of the isomorphism - EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) - EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) - - ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙ G (suc n))) → EM G n - ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) - - EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl - EM→ΩEM+1-0ₖ zero = emloop-1g _ - EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) - EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) - - EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) - → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y - EM→ΩEM+1-hom zero x y = emloop-comp _ x y - EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y - EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y - - EM→ΩEM+1-sym : (n : ℕ) (x : EM G n) → EM→ΩEM+1 n (-ₖ x) ≡ sym (EM→ΩEM+1 n x) - EM→ΩEM+1-sym n = - morphLemmas.distrMinus _+ₖ_ _∙_ (EM→ΩEM+1 n) (EM→ΩEM+1-hom n) - (0ₖ n) refl - -ₖ_ sym - (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) - (lCancelₖ n) rCancel - ∙assoc - (EM→ΩEM+1-0ₖ n) - - ΩEM+1→EM-sym : (n : ℕ) (p : typ (Ω (EM∙ G (suc n)))) - → ΩEM+1→EM n (sym p) ≡ -ₖ (ΩEM+1→EM n p) - ΩEM+1→EM-sym n p = sym (cong (ΩEM+1→EM n) (EM→ΩEM+1-sym n (ΩEM+1→EM n p) - ∙ cong sym (Iso.rightInv (Iso-EM-ΩEM+1 n) p))) - ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (-ₖ ΩEM+1→EM n p) - - ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙ G (suc n)))) - → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) - ΩEM+1→EM-hom n = - morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) - (EM→ΩEM+1-hom n) (ΩEM+1→EM n) - (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - - ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n - ΩEM+1→EM-refl zero = transportRefl 0g - ΩEM+1→EM-refl (suc zero) = refl - ΩEM+1→EM-refl (suc (suc n)) = refl - - EM→ΩEM+1∙ : (n : ℕ) → EM∙ G n →∙ Ω (EM∙ G (suc n)) - EM→ΩEM+1∙ n .fst = EM→ΩEM+1 n - EM→ΩEM+1∙ zero .snd = emloop-1g (AbGroup→Group G) - EM→ΩEM+1∙ (suc zero) .snd = cong (cong ∣_∣ₕ) (rCancel (merid embase)) - EM→ΩEM+1∙ (suc (suc n)) .snd = cong (cong ∣_∣ₕ) (rCancel (merid north)) - - EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) - EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) - - isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) - isHomogeneousEM n x = - ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) - - isCommΩEM-base : (n : ℕ) (x : _) - (p q : typ (Ω (EM G (suc n) , x))) → p ∙ q ≡ q ∙ p - isCommΩEM-base n = - EM-raw'-elim _ _ (λ _ → isOfHLevelΠ2 (2 + n) - λ _ _ → isOfHLevelPath (2 + n) - (hLevelEM _ (suc n) _ _) _ _) - (EM-raw'-trivElim _ _ - (λ _ → isOfHLevelΠ2 (suc n) λ _ _ → hLevelEM _ (suc n) _ _ _ _) - (lem n)) - where - * : (n : ℕ) → _ - * n = EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) - - lem : (n : ℕ) (p q : typ (Ω (EM G (suc n) , * n))) - → p ∙ q ≡ q ∙ p - lem zero = isCommΩEM zero - lem (suc n) = isCommΩEM (suc n) - - -- ΩEM+1→EM for arbitrarily based loops. Defining it by pattern - -- matching is more involved but should give better computational - -- properties. - ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) → x ≡ x → EM G n - ΩEM+1→EM-gen zero = - elimSet _ (λ _ → isSetΠ λ _ → is-set) (ΩEM+1→EM 0) - λ g → toPathP (funExt - λ q → transportRefl (ΩEM+1→EM 0 - (transport (λ i → Path (EM G (suc zero)) - (emloop g (~ i)) (emloop g (~ i))) q)) - ∙ cong (ΩEM+1→EM 0) - (fromPathP - (doubleCompPath-filler (emloop g) q (sym (emloop g)) - ▷ (doubleCompPath≡compPath _ _ _ - ∙∙ cong (emloop g ∙_) (isCommΩEM 0 q (sym (emloop g))) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ q) (rCancel (emloop g)) - ∙∙ sym (lUnit q))))) - ΩEM+1→EM-gen (suc n) = - Trunc.elim (λ _ → isOfHLevelΠ (4 + n) - λ _ → isOfHLevelSuc (3 + n) (hLevelEM _ (suc n))) - λ { north → ΩEM+1→EM (suc n) - ; south p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) - ∙∙ p - ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) - ; (merid a i) → help a i} - where - help : (a : EM-raw G (suc n)) - → PathP (λ i → Path (EM G (suc (suc n))) ∣ merid a i ∣ ∣ merid a i ∣ - → EM G (suc n)) - (ΩEM+1→EM (suc n)) - λ p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) - ∙∙ p - ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) - help a = - toPathP (funExt (λ p → - (transportRefl (ΩEM+1→EM (suc n) - (transport (λ i → Path (EM G (suc (suc n))) - ∣ merid a (~ i) ∣ ∣ merid a (~ i) ∣) p)) - ∙ cong (ΩEM+1→EM (suc n)) - (flipTransport - (((rUnit p - ∙ cong (p ∙_) (sym (lCancel _))) - ∙ isCommΩEM-base (suc n) ∣ south ∣ p _) - ∙∙ sym (∙assoc _ _ p) - ∙∙ cong₂ _∙_ (cong (cong ∣_∣ₕ) - (sym (cong sym (symDistr - (sym (merid a)) (merid ptEM-raw))))) - (isCommΩEM-base _ _ _ p) - ∙∙ sym (doubleCompPath≡compPath _ _ _) - ∙∙ λ j → transp (λ i → Path (EM G (suc (suc n))) - ∣ merid a (i ∨ ~ j) ∣ ∣ merid a (i ∨ ~ j) ∣) (~ j) - (cong ∣_∣ₕ (compPath-filler' - (sym (merid a)) (merid ptEM-raw) (~ j)) - ∙∙ p - ∙∙ cong ∣_∣ₕ (compPath-filler - (sym (merid ptEM-raw)) (merid a) (~ j)))))))) - - EM→ΩEM+1∘EM-raw→EM : (n : ℕ) (x : EM-raw G (suc n)) - → EM→ΩEM+1 (suc n) (EM-raw→EM _ _ x) ≡ cong ∣_∣ₕ (merid x ∙ sym (merid ptEM-raw)) - EM→ΩEM+1∘EM-raw→EM zero x = refl - EM→ΩEM+1∘EM-raw→EM (suc n) x = refl - - EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) - → EM G n → x ≡ x - EM→ΩEM+1-gen n x p = - sym (rUnitₖ (suc n) x) - ∙∙ cong (x +ₖ_) (EM→ΩEM+1 n p) - ∙∙ rUnitₖ (suc n) x - - ΩEM+1-gen→EM-0ₖ : (n : ℕ) (x : _) - → ΩEM+1→EM-gen n (0ₖ (suc n)) x - ≡ ΩEM+1→EM n x - ΩEM+1-gen→EM-0ₖ zero p = refl - ΩEM+1-gen→EM-0ₖ (suc n) p = refl - - EM→ΩEM+1-gen-0ₖ : (n : ℕ) (x : _) - → EM→ΩEM+1-gen n (0ₖ (suc n)) x - ≡ EM→ΩEM+1 n x - EM→ΩEM+1-gen-0ₖ zero x = sym (rUnit _) - ∙ λ j i → lUnitₖ 1 (EM→ΩEM+1 zero x i) j - EM→ΩEM+1-gen-0ₖ (suc n) x = sym (rUnit _) - ∙ λ j i → lUnitₖ (suc (suc n)) (EM→ΩEM+1 (suc n) x i) j - - EM→ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) - → (y : EM G n) → ΩEM+1→EM-gen n x (EM→ΩEM+1-gen n x y) ≡ y - EM→ΩEM+1→EM-gen n = - EM-raw'-elim _ _ - (λ _ → isOfHLevelΠ (suc (suc n)) - (λ _ → isOfHLevelPath (suc (suc n)) - (hLevelEM _ n) _ _)) - (EM-raw'-trivElim _ n - (λ _ → isOfHLevelΠ (suc n) λ _ → hLevelEM _ n _ _) - λ y → cong (λ p → ΩEM+1→EM-gen n p - (EM→ΩEM+1-gen n p y)) - (EM-raw'→EM∙ G (suc n)) - ∙ (λ i → ΩEM+1-gen→EM-0ₖ n (EM→ΩEM+1-gen-0ₖ n y i) i) - ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) y) - - ΩEM+1→EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) - → (y : x ≡ x) → EM→ΩEM+1-gen n x (ΩEM+1→EM-gen n x y) ≡ y - ΩEM+1→EM→ΩEM+1-gen n = - EM-raw'-elim _ _ - (λ _ → isOfHLevelΠ (suc (suc n)) - (λ _ → isOfHLevelPath (suc (suc n)) - (hLevelEM _ (suc n) _ _) _ _)) - (EM-raw'-trivElim _ n - (λ _ → isOfHLevelΠ (suc n) - λ _ → hLevelEM _ (suc n) _ _ _ _) - (subst (λ p → (y : p ≡ p) - → EM→ΩEM+1-gen n p (ΩEM+1→EM-gen n p y) ≡ y) - (sym (EM-raw'→EM∙ _ (suc n))) - λ p → (λ i → EM→ΩEM+1-gen-0ₖ n (ΩEM+1-gen→EM-0ₖ n p i) i) - ∙ Iso.rightInv (Iso-EM-ΩEM+1 n) p)) - - Iso-EM-ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) - → Iso (EM G n) (x ≡ x) - Iso.fun (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1-gen n x - Iso.inv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM-gen n x - Iso.rightInv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM→ΩEM+1-gen n x - Iso.leftInv (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1→EM-gen n x - - ΩEM+1→EM-gen-refl : (n : ℕ) (x : EM G (suc n)) - → ΩEM+1→EM-gen n x refl ≡ 0ₖ n - ΩEM+1→EM-gen-refl n = - EM-raw'-elim _ (suc n) - (λ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _) - (EM-raw'-trivElim _ n - (λ _ → hLevelEM _ n _ _) - (lem n)) - where - lem : (n : ℕ) → ΩEM+1→EM-gen n - (EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) refl - ≡ 0ₖ n - lem zero = ΩEM+1→EM-refl 0 - lem (suc n) = ΩEM+1→EM-refl (suc n) - - ΩEM+1→EM-gen-hom : (n : ℕ) (x : EM G (suc n)) (p q : x ≡ x) - → ΩEM+1→EM-gen n x (p ∙ q) ≡ ΩEM+1→EM-gen n x p +ₖ ΩEM+1→EM-gen n x q - ΩEM+1→EM-gen-hom n = - EM-raw'-elim _ (suc n) - (λ _ → isOfHLevelΠ2 (suc (suc n)) - (λ _ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _)) - (EM-raw'-trivElim _ n - (λ _ → isOfHLevelΠ2 (suc n) (λ _ _ → hLevelEM _ n _ _)) - (lem n)) - where - lem : (n : ℕ) (p q : EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) - ≡ EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) - → ΩEM+1→EM-gen n _ (p ∙ q) - ≡ ΩEM+1→EM-gen n _ p - +ₖ ΩEM+1→EM-gen n _ q - lem zero p q = ΩEM+1→EM-hom 0 p q - lem (suc n) p q = ΩEM+1→EM-hom (suc n) p q - --- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for --- the cup product -module _ where - open AbGroupStr renaming (_+_ to comp) - - isContr-↓∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ H n) - fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl - snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = - Σ≡Prop (λ x → is-set (snd H) _ _) - (funExt (raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) - fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl - fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = - Trunc.elim {B = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) - (raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) - x i - snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - - isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) - → isContr ((EM-raw G (suc n) , ptEM-raw) →∙ EM∙ H n) - isContr-↓∙' zero = isContr-↓∙ zero - fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl - fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = - raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i - snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) - - isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) - → isOfHLevel (suc m) (A →∙ EM∙ G n) - → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) - isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ - where - step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) - step₁ = - subst (isOfHLevel (suc m)) - (λ i → A →∙ ua∙ {A = Ω (EM∙ G (suc n))} {B = EM∙ G n} - (invEquiv (EM≃ΩEM+1 n)) - (ΩEM+1→EM-refl n) (~ i)) hlev - - step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ EM∙ G (suc n) ∙))) - step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ - - step₃ : isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) - step₃ = isOfHLevelΩ→isOfHLevel m - λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) - (isHomogeneous→∙ (isHomogeneousEM (suc n)) f) - step₂ - - isOfHLevel↑∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n m → isOfHLevel (2 + n) (EM∙ G (suc m) →∙ EM∙ H (suc (n + m))) - isOfHLevel↑∙ zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙ m)) - isOfHLevel↑∙ (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) - - isOfHLevel↑∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n m → isOfHLevel (2 + n) (EM-raw∙ G (suc m) →∙ EM∙ H (suc (n + m))) - isOfHLevel↑∙-lem zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) - isOfHLevel↑∙-lem (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙-lem n m) - - EM₁→∙Iso : {G : AbGroup ℓ} {H : AbGroup ℓ'} (m : ℕ) - → Iso (EM-raw'∙ G 1 →∙ EM∙ H (suc m)) (fst G → typ (Ω (EM∙ H (suc m)))) - Iso.fun (EM₁→∙Iso m) f g = sym (snd f) ∙∙ cong (fst f) (emloop-raw g) ∙∙ snd f - fst (Iso.inv (EM₁→∙Iso m) f) embase-raw = 0ₖ (suc m) - fst (Iso.inv (EM₁→∙Iso m) f) (emloop-raw g i) = f g i - snd (Iso.inv (EM₁→∙Iso m) f) = refl - Iso.rightInv (EM₁→∙Iso m) f = funExt λ x → sym (rUnit _) - Iso.leftInv (EM₁→∙Iso m) (f , p) = - →∙Homogeneous≡ (isHomogeneousEM _) - (funExt λ { embase-raw → sym p - ; (emloop-raw g i) j - → doubleCompPath-filler (sym p) (cong f (emloop-raw g)) p (~ j) i}) - - isOfHLevel↑∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n m → isOfHLevel (2 + n) (EM-raw'∙ G m →∙ EM∙ H (n + m)) - isOfHLevel↑∙' {H = H} zero zero = - isOfHLevelΣ 2 (isOfHLevelΠ 2 (λ _ → AbGroupStr.is-set (snd H))) - λ _ → isOfHLevelPath 2 (AbGroupStr.is-set (snd H)) _ _ - isOfHLevel↑∙' zero (suc zero) = - subst (isOfHLevel 2) (sym (isoToPath (EM₁→∙Iso 0))) - (isSetΠ λ _ → emsquash _ _) - isOfHLevel↑∙' zero (suc (suc m)) = isOfHLevel↑∙-lem zero (suc m) - isOfHLevel↑∙' {H = H} (suc n) zero = - isOfHLevelΣ (2 + suc n) (isOfHLevelΠ (2 + suc n) - (λ _ → subst (isOfHLevel (suc (suc (suc n)))) - (cong (EM H) (cong suc (+-comm 0 n))) - (hLevelEM H (suc n)))) - λ _ → isOfHLevelPath (suc (suc (suc n))) - (subst (isOfHLevel (suc (suc (suc n)))) - (cong (EM H) (cong suc (+-comm 0 n))) - (hLevelEM H (suc n))) _ _ - isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc zero) = - subst (isOfHLevel (2 + suc n)) (sym (isoToPath (EM₁→∙Iso (suc n))) - ∙ λ i → EM-raw'∙ G 1 →∙ EM∙ H (suc (+-comm 1 n i))) - (isOfHLevelΠ (2 + suc n) λ x → (isOfHLevelTrunc (4 + n) _ _)) - isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc (suc m)) = - subst (isOfHLevel (2 + suc n)) - (λ i → (EM-raw'∙ G (suc (suc m)) - →∙ EM∙ H (suc (+-suc n (suc m) (~ i))))) - (isOfHLevel↑∙-lem (suc n) (suc m)) - - →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) - → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) - →∙EMPath {G = G} A n = - ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl - ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) - - private - contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) - → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) - fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl - snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = - →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) - (sym (funExt (help n f p))) - where - help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) - → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) - → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help zero f p = - raw-elim G zero - (λ _ → isOfHLevel↑∙ zero m _ _) p - help (suc n) f p = - Trunc.elim (λ _ → isOfHLevelPath (4 + n) - (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) - (λ i → suc (suc (+-comm (suc n) 1 i))) - (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) - (raw-elim G (suc n) - (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) - - contr∙-lem' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) - → isContr (EM∙ G (suc n) →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) - fst (contr∙-lem' n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl - snd (contr∙-lem' {G = G} {H = H} {L = L} n m) (f , p) = - →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) - (funExt λ x → sym (help' n f p x)) - where - help' : (n : ℕ) → (f : EM G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) - → f (snd (EM∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) - → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help' zero f p = - raw-elim _ zero (λ _ → isOfHLevel↑∙' zero (suc m) _ _) p - help' (suc n) f p = - Trunc.elim (λ _ → isOfHLevelPath (4 + n) - (subst2 (λ x y → isOfHLevel x (EM-raw'∙ H (suc m) →∙ EM∙ L y)) - (λ i → suc (suc (suc (+-comm n 1 i)))) - (cong suc (+-suc n m)) - (isOfHLevelPlus' {n = 1} (suc (suc (suc n))) - (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m)))) _ _) - (raw-elim _ (suc n) (λ _ → isOfHLevelPath' (2 + n) - (subst (λ y → isOfHLevel (suc (suc (suc n))) - (EM-raw'∙ H (suc m) →∙ EM∙ L y)) - (+-suc (suc n) m) - (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p) - - contr∙-lem'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) - → isContr (EM-raw'∙ G (suc n) - →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) - fst (contr∙-lem'' n m) = (λ _ → (λ _ → 0ₖ (suc (n + m))) , refl) , refl - snd (contr∙-lem'' {G = G} {H = H} {L = L} n m) (f , p) = - →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) - (funExt λ x → sym (help n f p x)) - where - help : (n : ℕ) → (f : EM-raw' G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) - → f (snd (EM-raw'∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) - → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help zero f p = - EM-raw'-trivElim G zero (λ _ → isOfHLevel↑∙' _ _ _ _) p - help (suc n) f p = - EM-raw'-trivElim _ _ (λ _ → isOfHLevelPath' (suc (suc n)) - (subst (λ y → isOfHLevel (suc (suc (suc n))) (EM-raw'∙ H (suc m) →∙ EM∙ L y)) - (cong suc (+-suc n m)) - (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p - - isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) - →∙ (EM∙ H (suc m) - →∙ EM∙ L (suc (suc (l + n + m))) ∙)) - isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m zero = - isOfHLevelΩ→isOfHLevel 0 λ f - → subst - isProp (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) - where - h : isProp (EM∙ G (suc n) - →∙ (Ω (EM∙ H (suc m) - →∙ EM∙ L (suc (suc (n + m))) ∙))) - h = - subst isProp - (λ i → EM∙ G (suc n) - →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) - (isContr→isProp (contr∙-lem n m)) - isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = - isOfHLevelΩ→isOfHLevel (suc l) - λ f → - subst (isOfHLevel (2 + l)) - (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) - where - h : isOfHLevel (2 + l) - (EM∙ G (suc n) - →∙ (Ω (EM∙ H (suc m) - →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) - h = - subst (isOfHLevel (2 + l)) - (λ i → EM∙ G (suc n) - →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) - (isOfHLevel↑∙∙ n m l) - - isOfHLevel↑∙∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) - →∙ (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (l + n + m))) ∙)) - isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m zero = - isOfHLevelΩ→isOfHLevel 0 - λ f → subst - isProp (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) - lem) - where - lem : isProp (EM∙ G (suc n) - →∙ (Ω (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (n + m))) ∙))) - lem = subst isProp - (λ i → EM∙ G (suc n) - →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) - (isContr→isProp (contr∙-lem' n m)) - isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m (suc l) = - isOfHLevelΩ→isOfHLevel (suc l) - λ f → - subst (isOfHLevel (2 + l)) - (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) - where - lem : isOfHLevel (2 + l) - (EM∙ G (suc n) - →∙ (Ω (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) - lem = - subst (isOfHLevel (2 + l)) - (λ i → EM∙ G (suc n) - →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) - (isOfHLevel↑∙∙' n m l) - - isOfHLevel↑∙∙'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n m l → isOfHLevel (2 + l) (EM-raw'∙ G (suc n) - →∙ (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (l + n + m))) ∙)) - isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m zero = - isOfHLevelΩ→isOfHLevel 0 - λ f → subst - isProp (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) - lem) - where - lem : isProp (EM-raw'∙ G (suc n) - →∙ (Ω (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (n + m))) ∙))) - lem = subst isProp - (λ i → EM-raw'∙ G (suc n) - →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) - (isContr→isProp (contr∙-lem'' _ _)) - isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m (suc l) = - isOfHLevelΩ→isOfHLevel (suc l) - λ f → - subst (isOfHLevel (2 + l)) - (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) - (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) - where - lem : isOfHLevel (2 + l) - (EM-raw'∙ G (suc n) - →∙ (Ω (EM-raw'∙ H (suc m) - →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) - lem = - subst (isOfHLevel (2 + l)) - (λ i → EM-raw'∙ G (suc n) - →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) - (isOfHLevel↑∙∙'' n m l) - --- A homomorphism φ : G → H of AbGroups induces a homomorphism --- φ' : K(G,n) → K(H,n) -inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupHom G' H' - → ∀ n - → EM-raw G' n → EM-raw H' n -inducedFun-EM-raw f = - elim+2 (fst f) - (EMrec _ emsquash embase - (λ g → emloop (fst f g)) - λ g h → compPathR→PathP (sym - (sym (lUnit _) - ∙∙ cong (_∙ (sym (emloop (fst f h)))) - (cong emloop (IsGroupHom.pres· (snd f) g h) - ∙ emloop-comp _ (fst f g) (fst f h)) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) - ∙∙ sym (rUnit _)))) - (λ n ind → λ { north → north - ; south → south - ; (merid a i) → merid (ind a) i} ) - -inducedFun-EM-raw-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM-raw G' n) - → inducedFun-EM-raw (idGroupHom {G = AbGroup→Group G'}) n x ≡ x -inducedFun-EM-raw-id zero x = refl -inducedFun-EM-raw-id (suc zero) = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) - λ { embase-raw → refl ; (emloop-raw g i) → refl} -inducedFun-EM-raw-id (suc (suc n)) north = refl -inducedFun-EM-raw-id (suc (suc n)) south = refl -inducedFun-EM-raw-id (suc (suc n)) (merid a i) j = merid (inducedFun-EM-raw-id (suc n) a j) i - -inducedFun-EM-raw-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} - (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) - → (x : EM-raw G' n) → inducedFun-EM-raw (compGroupHom ϕ ψ) n x - ≡ inducedFun-EM-raw ψ n (inducedFun-EM-raw ϕ n x) -inducedFun-EM-raw-comp ϕ ψ zero x = refl -inducedFun-EM-raw-comp ϕ ψ (suc zero) = - EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) - λ { embase-raw → refl ; (emloop-raw g i) → refl} -inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) north = refl -inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) south = refl -inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) (merid a i) j = - merid (inducedFun-EM-raw-comp ϕ ψ (suc n) a j) i - -inducedFun-EM : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupHom G' H' - → ∀ n - → EM G' n → EM H' n -inducedFun-EM f zero = inducedFun-EM-raw f zero -inducedFun-EM f (suc zero) = inducedFun-EM-raw f (suc zero) -inducedFun-EM f (suc (suc n)) = Trunc.map (inducedFun-EM-raw f (2 + n)) - -EM-raw→EM-funct : {G : AbGroup ℓ} {H : AbGroup ℓ'} - (n : ℕ) (ψ : AbGroupHom G H) (y : EM-raw G n) - → EM-raw→EM _ _ (inducedFun-EM-raw ψ n y) - ≡ inducedFun-EM ψ n (EM-raw→EM _ _ y) -EM-raw→EM-funct zero ψ y = refl -EM-raw→EM-funct (suc zero) ψ y = refl -EM-raw→EM-funct (suc (suc n)) ψ y = refl - -inducedFun-EM-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM G' n) - → inducedFun-EM (idGroupHom {G = AbGroup→Group G'}) n x ≡ x -inducedFun-EM-id zero x = refl -inducedFun-EM-id (suc zero) = inducedFun-EM-raw-id (suc zero) -inducedFun-EM-id (suc (suc n)) = - Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) - λ x → cong ∣_∣ₕ (inducedFun-EM-raw-id _ x) - -inducedFun-EM-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} - (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) - → (x : EM G' n) → inducedFun-EM (compGroupHom ϕ ψ) n x - ≡ inducedFun-EM ψ n (inducedFun-EM ϕ n x) -inducedFun-EM-comp ϕ ψ zero x = refl -inducedFun-EM-comp ϕ ψ (suc zero) = inducedFun-EM-raw-comp ϕ ψ (suc zero) -inducedFun-EM-comp ϕ ψ (suc (suc n)) = - Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) - λ x → cong ∣_∣ₕ (inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) x) - -inducedFun-EM0ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {ϕ : AbGroupHom G' H'} (n : ℕ) - → inducedFun-EM ϕ n (0ₖ n) ≡ 0ₖ n -inducedFun-EM0ₖ {ϕ = ϕ} zero = IsGroupHom.pres1 (snd ϕ) -inducedFun-EM0ₖ (suc zero) = refl -inducedFun-EM0ₖ (suc (suc n)) = refl - -inducedFun-EM-pres+ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - (ϕ : AbGroupHom G' H') (n : ℕ) (x y : EM G' n) - → inducedFun-EM ϕ n (x +ₖ y) ≡ inducedFun-EM ϕ n x +ₖ inducedFun-EM ϕ n y -inducedFun-EM-pres+ₖ ϕ zero x y = IsGroupHom.pres· (snd ϕ) x y -inducedFun-EM-pres+ₖ {G' = G'} {H' = H'} ϕ (suc n) = - EM-elim2 (suc n) (λ _ _ → isOfHLevelPath (2 + suc n) (hLevelEM _ (suc n)) _ _) - (wedgeConEM.fun _ _ n n - (λ _ _ → isOfHLevelPath' (suc n + suc n) - (subst (λ m → isOfHLevel (suc (suc m)) (EM H' (suc n))) - (sym (+-suc n n)) - (isOfHLevelPlus' {n = n} (3 + n) - (hLevelEM _ (suc n)))) - _ _) - (l n) - (r n) - (l≡r n)) - where - lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) → EM-raw→EM G (suc n) ptEM-raw ≡ 0ₖ _ - lem zero = refl - lem (suc n) = refl - - l : (n : ℕ) (y : EM-raw G' (suc n)) - → inducedFun-EM ϕ (suc n) ((EM-raw→EM G' (suc n) ptEM-raw) - +ₖ EM-raw→EM G' (suc n) y) - ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) - +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y) - l n y = (cong (inducedFun-EM ϕ (suc n)) - (cong (_+ₖ EM-raw→EM G' (suc n) y) (lem n) - ∙ lUnitₖ (suc n) (EM-raw→EM G' (suc n) y))) - ∙∙ sym (lUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) - ∙∙ cong (_+ₖ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) - (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) - ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) - - r : (n : ℕ) (x : EM-raw G' (suc n)) - → inducedFun-EM ϕ (suc n) - (EM-raw→EM G' (suc n) x +ₖ EM-raw→EM G' (suc n) ptEM-raw) - ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) - +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) - r n x = cong (inducedFun-EM ϕ (suc n)) - (cong (EM-raw→EM G' (suc n) x +ₖ_) (lem n) - ∙ rUnitₖ (suc n) (EM-raw→EM G' (suc n) x)) - ∙∙ sym (rUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x))) - ∙∙ cong (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) +ₖ_) - (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) - ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) - - l≡r : (n : ℕ) → l n ptEM-raw ≡ r n ptEM-raw - l≡r zero = refl - l≡r (suc n) = refl - -inducedFun-EM-pres-ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - (ϕ : AbGroupHom G' H') (n : ℕ) (x : EM G' n) - → inducedFun-EM ϕ n (-ₖ x) ≡ -ₖ (inducedFun-EM ϕ n x) -inducedFun-EM-pres-ₖ ϕ n = - morphLemmas.distrMinus _+ₖ_ _+ₖ_ - (inducedFun-EM ϕ n) (inducedFun-EM-pres+ₖ ϕ n) (0ₖ n) (0ₖ n) (-ₖ_) (-ₖ_) - (lUnitₖ n) (rUnitₖ n) - (lCancelₖ n) (rCancelₖ n) - (assocₖ n) (inducedFun-EM0ₖ n) - -EMFun-EM→ΩEM+1 : {G : AbGroup ℓ} {H : AbGroup ℓ'} - {ϕ : AbGroupHom G H} (n : ℕ) (x : EM G n) - → PathP (λ i → inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i) - ≡ inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i)) - (EM→ΩEM+1 n (inducedFun-EM ϕ n x)) - (cong (inducedFun-EM ϕ (suc n)) (EM→ΩEM+1 n x)) -EMFun-EM→ΩEM+1 {ϕ = ϕ} zero x = refl -EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc zero) x = - cong-∙ ∣_∣ₕ (merid (inducedFun-EM ϕ (suc zero) x)) - (sym (merid embase)) - ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc zero))) - (cong ∣_∣ₕ (merid x)) (cong ∣_∣ₕ (sym (merid embase)))) - ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc zero)))) - (sym (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase)))) -EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc (suc n)) = - Trunc.elim (λ _ → isOfHLevelPath (4 + n) - (isOfHLevelTrunc (5 + n) _ _) _ _) - λ a → cong-∙ ∣_∣ₕ (merid (inducedFun-EM-raw ϕ (2 + n) a)) - (sym (merid north)) - ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc (suc n)))) - (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (sym (merid north)))) - ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc (suc n))))) - (sym (cong-∙ ∣_∣ₕ (merid a) (sym (merid north)))) - - -inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupEquiv G' H' - → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) -Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n -Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n -Iso.rightInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) - (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) - h = elim+2 - (secEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - (sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) - (sym (lUnit _)) - ∙ rCancel _)))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} -Iso.leftInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) - (Iso.inv (inducedFun-EM-rawIso e n)) - h = elim+2 - (retEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) - ∙ rCancel _))))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} - -module _ {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) where - Iso→EMIso : ∀ n → Iso (EM G n) (EM H n) - Iso.fun (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom e) n - Iso.inv (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom (invGroupEquiv e)) n - Iso.rightInv (Iso→EMIso zero) = Iso.rightInv (inducedFun-EM-rawIso e zero) - Iso.rightInv (Iso→EMIso (suc zero)) = - Iso.rightInv (inducedFun-EM-rawIso e (suc zero)) - Iso.rightInv (Iso→EMIso (suc (suc n))) = - Iso.rightInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) - Iso.leftInv (Iso→EMIso zero) = - Iso.leftInv (inducedFun-EM-rawIso e zero) - Iso.leftInv (Iso→EMIso (suc zero)) = - Iso.leftInv (inducedFun-EM-rawIso e (suc zero)) - Iso.leftInv (Iso→EMIso (suc (suc n))) = - Iso.leftInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) - - Iso→EMIso∙ : ∀ n → Iso.fun (Iso→EMIso n) (EM∙ G n .snd) ≡ EM∙ H n .snd - Iso→EMIso∙ zero = IsGroupHom.pres1 (e .snd) - Iso→EMIso∙ (suc zero) = refl - Iso→EMIso∙ (suc (suc n)) = refl - - Iso→EMIso⁻∙ : ∀ n → Iso.inv (Iso→EMIso n) (EM∙ H n .snd) ≡ EM∙ G n .snd - Iso→EMIso⁻∙ zero = IsGroupHom.pres1 (invGroupEquiv e .snd) - Iso→EMIso⁻∙ (suc zero) = refl - Iso→EMIso⁻∙ (suc (suc n)) = refl - -Iso→EMIsoInv : {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) - → ∀ n → Iso.inv (Iso→EMIso e n) ≡ Iso.fun (Iso→EMIso (invGroupEquiv e) n) -Iso→EMIsoInv e zero = refl -Iso→EMIsoInv e (suc zero) = refl -Iso→EMIsoInv e (suc (suc n)) = refl - -EM⊗-commIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) -EM⊗-commIso {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) - -EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) -EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) +-- isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) +-- isGroupoidEM₁ = emsquash + +-- --------- Ω (EM₁ G) ≃ G --------- + +-- {- since we write composition in diagrammatic order, +-- and function composition in the other order, +-- we need right multiplication here -} +-- rightEquiv : (g : G) → G ≃ G +-- rightEquiv g = isoToEquiv isom +-- where +-- isom : Iso G G +-- isom .Iso.fun = _· g +-- isom .Iso.inv = _· inv g +-- isom .Iso.rightInv h = +-- (h · inv g) · g ≡⟨ (·Assoc h (inv g) g) ⁻¹ ⟩ +-- h · inv g · g ≡⟨ cong (h ·_) (·InvL g) ⟩ +-- h · 1g ≡⟨ ·IdR h ⟩ h ∎ +-- isom .Iso.leftInv h = +-- (h · g) · inv g ≡⟨ (·Assoc h g (inv g)) ⁻¹ ⟩ +-- h · g · inv g ≡⟨ cong (h ·_) (·InvR g) ⟩ +-- h · 1g ≡⟨ ·IdR h ⟩ h ∎ + +-- compRightEquiv : (g h : G) +-- → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) +-- compRightEquiv g h = equivEq (funExt (λ x → (·Assoc x g h) ⁻¹)) + +-- CodesSet : EM₁ Ĝ → hSet ℓ +-- CodesSet = EMrec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp +-- where +-- RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) +-- RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) + +-- lemma₁ : (g h : G) → Square +-- (ua (rightEquiv g)) (ua (rightEquiv (g · h))) +-- refl (ua (rightEquiv h)) +-- lemma₁ g h = invEq +-- (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) +-- refl (ua (rightEquiv h))) +-- (ua (rightEquiv g) ∙ ua (rightEquiv h) +-- ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ +-- ua (compEquiv (rightEquiv g) (rightEquiv h)) +-- ≡⟨ cong ua (compRightEquiv g h) ⟩ +-- ua (rightEquiv (g · h)) ∎) + +-- REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) +-- REComp g h = ΣSquareSet (λ _ → isProp→isSet isPropIsSet) (lemma₁ g h) +-- Codes : EM₁ Ĝ → Type ℓ +-- Codes x = CodesSet x .fst + +-- encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x +-- encode x p = subst Codes p 1g + +-- decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x +-- decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) +-- emloop λ g → ua→ λ h → emcomp h g + +-- decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p +-- decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) +-- (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ +-- emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p + +-- encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c +-- encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) +-- λ g → encode embase (decode embase g) ≡⟨ refl ⟩ +-- encode embase (emloop g) ≡⟨ refl ⟩ +-- transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ +-- 1g · g ≡⟨ ·IdL g ⟩ +-- g ∎ + +-- ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G +-- Iso.fun ΩEM₁Iso = encode embase +-- Iso.inv ΩEM₁Iso = emloop +-- Iso.rightInv ΩEM₁Iso = encode-decode embase +-- Iso.leftInv ΩEM₁Iso = decode-encode embase + +-- ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G +-- ΩEM₁≡ = isoToPath ΩEM₁Iso + +-- --------- Ω (EMₙ₊₁ G) ≃ EMₙ G --------- +-- module _ {G : AbGroup ℓ} where +-- open AbGroupStr (snd G) +-- renaming (_+_ to _+G_ ; -_ to -G_ ; +Assoc to +AssocG) + +-- CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓ (3 + n) +-- CODE n = +-- Trunc.elim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) +-- λ { north → EM G (suc n) , hLevelEM G (suc n) +-- ; south → EM G (suc n) , hLevelEM G (suc n) +-- ; (merid a i) → fib n a i} +-- where +-- fib : (n : ℕ) → (a : EM-raw G (suc n)) +-- → Path (TypeOfHLevel _ (3 + n)) +-- (EM G (suc n) , hLevelEM G (suc n)) +-- (EM G (suc n) , hLevelEM G (suc n)) +-- fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) +-- (isoToPath (addIso 1 a)) +-- fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) +-- (isoToPath (addIso (suc (suc n)) ∣ a ∣)) + +-- decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x +-- decode' n = +-- Trunc.elim (λ _ → isOfHLevelΠ (4 + n) +-- λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) +-- λ { north → f n +-- ; south → g n +-- ; (merid a i) → main a i} +-- where +-- f : (n : ℕ) → _ +-- f n = σ-EM' n + +-- g : (n : ℕ) → EM G (suc n) → ∣ ptEM-raw ∣ ≡ ∣ south ∣ +-- g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptEM-raw) + +-- lem₂ : (n : ℕ) (a x : _) +-- → Path (Path (EM G (suc (suc n))) _ _) +-- ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) +-- (g n (EM-raw→EM G (suc n) x)) +-- lem₂ zero a x = +-- cong (f zero x ∙_) +-- (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) +-- ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) +-- ∙∙ sym (∙assoc _ _ _) +-- ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) +-- ∙ sym (rUnit _)) +-- lem₂ (suc n) a x = +-- cong (f (suc n) ∣ x ∣ ∙_) +-- ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) +-- ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) +-- ∙∙ sym (∙assoc _ _ _) +-- ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) +-- ∙ sym (rUnit _))) + +-- lem : (n : ℕ) (x a : EM-raw G (suc n)) +-- → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) +-- (EM-raw→EM G (suc n) x)) +-- ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) +-- lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) +-- ∙∙ σ-EM'-hom zero x (-ₖ a) +-- ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) +-- lem (suc n) x a = +-- cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) +-- ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) +-- ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) + +-- main : (a : _) +-- → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst +-- → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) +-- main a = +-- toPathP (funExt +-- (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) +-- λ x → +-- ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (i ∨ j) ∣) +-- i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) +-- ∙∙ sym (∙assoc _ _ _) +-- ∙∙ lem₂ n a x)) + +-- encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst +-- encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) + +-- decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) +-- → decode' n x (encode' n x p) ≡ p +-- decode'-encode' zero x = +-- J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) +-- (σ-EM'-0ₖ 0) +-- decode'-encode' (suc n) x = +-- J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) +-- (σ-EM'-0ₖ (suc n)) + +-- encode'-decode' : (n : ℕ) (x : _) +-- → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x +-- encode'-decode' zero x = +-- cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) +-- ∙∙ substComposite (λ x → CODE zero x .fst) +-- (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) embase +-- ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) +-- (λ i → lUnitₖ 1 (transportRefl x i) i) +-- ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) +-- encode'-decode' (suc n) = +-- Trunc.elim (λ _ → isOfHLevelTruncPath {n = 4 + n}) +-- λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) +-- ∙∙ substComposite (λ x → CODE (suc n) x .fst) +-- (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) (0ₖ (2 + n)) +-- ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptEM-raw (~ i) ∣ₕ)) +-- (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) +-- ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + +-- Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙ G (suc n)))) +-- Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) +-- Iso.fun (Iso-EM-ΩEM+1 (suc n)) = decode' n (0ₖ (2 + n)) +-- Iso.inv (Iso-EM-ΩEM+1 (suc n)) = encode' n ∣ north ∣ +-- Iso.rightInv (Iso-EM-ΩEM+1 (suc n)) = decode'-encode' _ _ +-- Iso.leftInv (Iso-EM-ΩEM+1 (suc n)) = encode'-decode' _ + +-- EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) +-- EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) + +-- -- Some properties of the isomorphism +-- EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) +-- EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) + +-- ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙ G (suc n))) → EM G n +-- ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + +-- EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl +-- EM→ΩEM+1-0ₖ zero = emloop-1g _ +-- EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) +-- EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) + +-- EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) +-- → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y +-- EM→ΩEM+1-hom zero x y = emloop-comp _ x y +-- EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y +-- EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y + +-- EM→ΩEM+1-sym : (n : ℕ) (x : EM G n) → EM→ΩEM+1 n (-ₖ x) ≡ sym (EM→ΩEM+1 n x) +-- EM→ΩEM+1-sym n = +-- morphLemmas.distrMinus _+ₖ_ _∙_ (EM→ΩEM+1 n) (EM→ΩEM+1-hom n) +-- (0ₖ n) refl +-- -ₖ_ sym +-- (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) +-- (lCancelₖ n) rCancel +-- ∙assoc +-- (EM→ΩEM+1-0ₖ n) + +-- ΩEM+1→EM-sym : (n : ℕ) (p : typ (Ω (EM∙ G (suc n)))) +-- → ΩEM+1→EM n (sym p) ≡ -ₖ (ΩEM+1→EM n p) +-- ΩEM+1→EM-sym n p = sym (cong (ΩEM+1→EM n) (EM→ΩEM+1-sym n (ΩEM+1→EM n p) +-- ∙ cong sym (Iso.rightInv (Iso-EM-ΩEM+1 n) p))) +-- ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (-ₖ ΩEM+1→EM n p) + +-- ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙ G (suc n)))) +-- → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) +-- ΩEM+1→EM-hom n = +-- morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) +-- (EM→ΩEM+1-hom n) (ΩEM+1→EM n) +-- (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) + +-- ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n +-- ΩEM+1→EM-refl zero = transportRefl 0g +-- ΩEM+1→EM-refl (suc zero) = refl +-- ΩEM+1→EM-refl (suc (suc n)) = refl + +-- EM→ΩEM+1∙ : (n : ℕ) → EM∙ G n →∙ Ω (EM∙ G (suc n)) +-- EM→ΩEM+1∙ n .fst = EM→ΩEM+1 n +-- EM→ΩEM+1∙ zero .snd = emloop-1g (AbGroup→Group G) +-- EM→ΩEM+1∙ (suc zero) .snd = cong (cong ∣_∣ₕ) (rCancel (merid embase)) +-- EM→ΩEM+1∙ (suc (suc n)) .snd = cong (cong ∣_∣ₕ) (rCancel (merid north)) + +-- EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) +-- EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) + +-- isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) +-- isHomogeneousEM n x = +-- ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) + +-- isCommΩEM-base : (n : ℕ) (x : _) +-- (p q : typ (Ω (EM G (suc n) , x))) → p ∙ q ≡ q ∙ p +-- isCommΩEM-base n = +-- EM-raw'-elim _ _ (λ _ → isOfHLevelΠ2 (2 + n) +-- λ _ _ → isOfHLevelPath (2 + n) +-- (hLevelEM _ (suc n) _ _) _ _) +-- (EM-raw'-trivElim _ _ +-- (λ _ → isOfHLevelΠ2 (suc n) λ _ _ → hLevelEM _ (suc n) _ _ _ _) +-- (lem n)) +-- where +-- * : (n : ℕ) → _ +-- * n = EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) + +-- lem : (n : ℕ) (p q : typ (Ω (EM G (suc n) , * n))) +-- → p ∙ q ≡ q ∙ p +-- lem zero = isCommΩEM zero +-- lem (suc n) = isCommΩEM (suc n) + +-- -- ΩEM+1→EM for arbitrarily based loops. Defining it by pattern +-- -- matching is more involved but should give better computational +-- -- properties. +-- ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) → x ≡ x → EM G n +-- ΩEM+1→EM-gen zero = +-- elimSet _ (λ _ → isSetΠ λ _ → is-set) (ΩEM+1→EM 0) +-- λ g → toPathP (funExt +-- λ q → transportRefl (ΩEM+1→EM 0 +-- (transport (λ i → Path (EM G (suc zero)) +-- (emloop g (~ i)) (emloop g (~ i))) q)) +-- ∙ cong (ΩEM+1→EM 0) +-- (fromPathP +-- (doubleCompPath-filler (emloop g) q (sym (emloop g)) +-- ▷ (doubleCompPath≡compPath _ _ _ +-- ∙∙ cong (emloop g ∙_) (isCommΩEM 0 q (sym (emloop g))) +-- ∙∙ ∙assoc _ _ _ +-- ∙∙ cong (_∙ q) (rCancel (emloop g)) +-- ∙∙ sym (lUnit q))))) +-- ΩEM+1→EM-gen (suc n) = +-- Trunc.elim (λ _ → isOfHLevelΠ (4 + n) +-- λ _ → isOfHLevelSuc (3 + n) (hLevelEM _ (suc n))) +-- λ { north → ΩEM+1→EM (suc n) +-- ; south p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) +-- ∙∙ p +-- ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) +-- ; (merid a i) → help a i} +-- where +-- help : (a : EM-raw G (suc n)) +-- → PathP (λ i → Path (EM G (suc (suc n))) ∣ merid a i ∣ ∣ merid a i ∣ +-- → EM G (suc n)) +-- (ΩEM+1→EM (suc n)) +-- λ p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) +-- ∙∙ p +-- ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) +-- help a = +-- toPathP (funExt (λ p → +-- (transportRefl (ΩEM+1→EM (suc n) +-- (transport (λ i → Path (EM G (suc (suc n))) +-- ∣ merid a (~ i) ∣ ∣ merid a (~ i) ∣) p)) +-- ∙ cong (ΩEM+1→EM (suc n)) +-- (flipTransport +-- (((rUnit p +-- ∙ cong (p ∙_) (sym (lCancel _))) +-- ∙ isCommΩEM-base (suc n) ∣ south ∣ p _) +-- ∙∙ sym (∙assoc _ _ p) +-- ∙∙ cong₂ _∙_ (cong (cong ∣_∣ₕ) +-- (sym (cong sym (symDistr +-- (sym (merid a)) (merid ptEM-raw))))) +-- (isCommΩEM-base _ _ _ p) +-- ∙∙ sym (doubleCompPath≡compPath _ _ _) +-- ∙∙ λ j → transp (λ i → Path (EM G (suc (suc n))) +-- ∣ merid a (i ∨ ~ j) ∣ ∣ merid a (i ∨ ~ j) ∣) (~ j) +-- (cong ∣_∣ₕ (compPath-filler' +-- (sym (merid a)) (merid ptEM-raw) (~ j)) +-- ∙∙ p +-- ∙∙ cong ∣_∣ₕ (compPath-filler +-- (sym (merid ptEM-raw)) (merid a) (~ j)))))))) + +-- EM→ΩEM+1∘EM-raw→EM : (n : ℕ) (x : EM-raw G (suc n)) +-- → EM→ΩEM+1 (suc n) (EM-raw→EM _ _ x) ≡ cong ∣_∣ₕ (merid x ∙ sym (merid ptEM-raw)) +-- EM→ΩEM+1∘EM-raw→EM zero x = refl +-- EM→ΩEM+1∘EM-raw→EM (suc n) x = refl + +-- EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) +-- → EM G n → x ≡ x +-- EM→ΩEM+1-gen n x p = +-- sym (rUnitₖ (suc n) x) +-- ∙∙ cong (x +ₖ_) (EM→ΩEM+1 n p) +-- ∙∙ rUnitₖ (suc n) x + +-- ΩEM+1-gen→EM-0ₖ : (n : ℕ) (x : _) +-- → ΩEM+1→EM-gen n (0ₖ (suc n)) x +-- ≡ ΩEM+1→EM n x +-- ΩEM+1-gen→EM-0ₖ zero p = refl +-- ΩEM+1-gen→EM-0ₖ (suc n) p = refl + +-- EM→ΩEM+1-gen-0ₖ : (n : ℕ) (x : _) +-- → EM→ΩEM+1-gen n (0ₖ (suc n)) x +-- ≡ EM→ΩEM+1 n x +-- EM→ΩEM+1-gen-0ₖ zero x = sym (rUnit _) +-- ∙ λ j i → lUnitₖ 1 (EM→ΩEM+1 zero x i) j +-- EM→ΩEM+1-gen-0ₖ (suc n) x = sym (rUnit _) +-- ∙ λ j i → lUnitₖ (suc (suc n)) (EM→ΩEM+1 (suc n) x i) j + +-- EM→ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) +-- → (y : EM G n) → ΩEM+1→EM-gen n x (EM→ΩEM+1-gen n x y) ≡ y +-- EM→ΩEM+1→EM-gen n = +-- EM-raw'-elim _ _ +-- (λ _ → isOfHLevelΠ (suc (suc n)) +-- (λ _ → isOfHLevelPath (suc (suc n)) +-- (hLevelEM _ n) _ _)) +-- (EM-raw'-trivElim _ n +-- (λ _ → isOfHLevelΠ (suc n) λ _ → hLevelEM _ n _ _) +-- λ y → cong (λ p → ΩEM+1→EM-gen n p +-- (EM→ΩEM+1-gen n p y)) +-- (EM-raw'→EM∙ G (suc n)) +-- ∙ (λ i → ΩEM+1-gen→EM-0ₖ n (EM→ΩEM+1-gen-0ₖ n y i) i) +-- ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) y) + +-- ΩEM+1→EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) +-- → (y : x ≡ x) → EM→ΩEM+1-gen n x (ΩEM+1→EM-gen n x y) ≡ y +-- ΩEM+1→EM→ΩEM+1-gen n = +-- EM-raw'-elim _ _ +-- (λ _ → isOfHLevelΠ (suc (suc n)) +-- (λ _ → isOfHLevelPath (suc (suc n)) +-- (hLevelEM _ (suc n) _ _) _ _)) +-- (EM-raw'-trivElim _ n +-- (λ _ → isOfHLevelΠ (suc n) +-- λ _ → hLevelEM _ (suc n) _ _ _ _) +-- (subst (λ p → (y : p ≡ p) +-- → EM→ΩEM+1-gen n p (ΩEM+1→EM-gen n p y) ≡ y) +-- (sym (EM-raw'→EM∙ _ (suc n))) +-- λ p → (λ i → EM→ΩEM+1-gen-0ₖ n (ΩEM+1-gen→EM-0ₖ n p i) i) +-- ∙ Iso.rightInv (Iso-EM-ΩEM+1 n) p)) + +-- Iso-EM-ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) +-- → Iso (EM G n) (x ≡ x) +-- Iso.fun (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1-gen n x +-- Iso.inv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM-gen n x +-- Iso.rightInv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM→ΩEM+1-gen n x +-- Iso.leftInv (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1→EM-gen n x + +-- ΩEM+1→EM-gen-refl : (n : ℕ) (x : EM G (suc n)) +-- → ΩEM+1→EM-gen n x refl ≡ 0ₖ n +-- ΩEM+1→EM-gen-refl n = +-- EM-raw'-elim _ (suc n) +-- (λ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _) +-- (EM-raw'-trivElim _ n +-- (λ _ → hLevelEM _ n _ _) +-- (lem n)) +-- where +-- lem : (n : ℕ) → ΩEM+1→EM-gen n +-- (EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) refl +-- ≡ 0ₖ n +-- lem zero = ΩEM+1→EM-refl 0 +-- lem (suc n) = ΩEM+1→EM-refl (suc n) + +-- ΩEM+1→EM-gen-hom : (n : ℕ) (x : EM G (suc n)) (p q : x ≡ x) +-- → ΩEM+1→EM-gen n x (p ∙ q) ≡ ΩEM+1→EM-gen n x p +ₖ ΩEM+1→EM-gen n x q +-- ΩEM+1→EM-gen-hom n = +-- EM-raw'-elim _ (suc n) +-- (λ _ → isOfHLevelΠ2 (suc (suc n)) +-- (λ _ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _)) +-- (EM-raw'-trivElim _ n +-- (λ _ → isOfHLevelΠ2 (suc n) (λ _ _ → hLevelEM _ n _ _)) +-- (lem n)) +-- where +-- lem : (n : ℕ) (p q : EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) +-- ≡ EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) +-- → ΩEM+1→EM-gen n _ (p ∙ q) +-- ≡ ΩEM+1→EM-gen n _ p +-- +ₖ ΩEM+1→EM-gen n _ q +-- lem zero p q = ΩEM+1→EM-hom 0 p q +-- lem (suc n) p q = ΩEM+1→EM-hom (suc n) p q + +-- -- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for +-- -- the cup product +-- module _ where +-- open AbGroupStr renaming (_+_ to comp) + +-- isContr-↓∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ H n) +-- fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl +-- snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = +-- Σ≡Prop (λ x → is-set (snd H) _ _) +-- (funExt (raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) +-- fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl +-- fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = +-- Trunc.elim {B = λ x → 0ₖ (suc n) ≡ fst f x} +-- (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) +-- (raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) +-- x i +-- snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +-- isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) +-- → isContr ((EM-raw G (suc n) , ptEM-raw) →∙ EM∙ H n) +-- isContr-↓∙' zero = isContr-↓∙ zero +-- fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl +-- fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = +-- raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} +-- (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i +-- snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +-- isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) +-- → isOfHLevel (suc m) (A →∙ EM∙ G n) +-- → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) +-- isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ +-- where +-- step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) +-- step₁ = +-- subst (isOfHLevel (suc m)) +-- (λ i → A →∙ ua∙ {A = Ω (EM∙ G (suc n))} {B = EM∙ G n} +-- (invEquiv (EM≃ΩEM+1 n)) +-- (ΩEM+1→EM-refl n) (~ i)) hlev + +-- step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ EM∙ G (suc n) ∙))) +-- step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + +-- step₃ : isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) +-- step₃ = isOfHLevelΩ→isOfHLevel m +-- λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) +-- (isHomogeneous→∙ (isHomogeneousEM (suc n)) f) +-- step₂ + +-- isOfHLevel↑∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- → ∀ n m → isOfHLevel (2 + n) (EM∙ G (suc m) →∙ EM∙ H (suc (n + m))) +-- isOfHLevel↑∙ zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙ m)) +-- isOfHLevel↑∙ (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +-- isOfHLevel↑∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- → ∀ n m → isOfHLevel (2 + n) (EM-raw∙ G (suc m) →∙ EM∙ H (suc (n + m))) +-- isOfHLevel↑∙-lem zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) +-- isOfHLevel↑∙-lem (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙-lem n m) + +-- EM₁→∙Iso : {G : AbGroup ℓ} {H : AbGroup ℓ'} (m : ℕ) +-- → Iso (EM-raw'∙ G 1 →∙ EM∙ H (suc m)) (fst G → typ (Ω (EM∙ H (suc m)))) +-- Iso.fun (EM₁→∙Iso m) f g = sym (snd f) ∙∙ cong (fst f) (emloop-raw g) ∙∙ snd f +-- fst (Iso.inv (EM₁→∙Iso m) f) embase-raw = 0ₖ (suc m) +-- fst (Iso.inv (EM₁→∙Iso m) f) (emloop-raw g i) = f g i +-- snd (Iso.inv (EM₁→∙Iso m) f) = refl +-- Iso.rightInv (EM₁→∙Iso m) f = funExt λ x → sym (rUnit _) +-- Iso.leftInv (EM₁→∙Iso m) (f , p) = +-- →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt λ { embase-raw → sym p +-- ; (emloop-raw g i) j +-- → doubleCompPath-filler (sym p) (cong f (emloop-raw g)) p (~ j) i}) + +-- isOfHLevel↑∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- → ∀ n m → isOfHLevel (2 + n) (EM-raw'∙ G m →∙ EM∙ H (n + m)) +-- isOfHLevel↑∙' {H = H} zero zero = +-- isOfHLevelΣ 2 (isOfHLevelΠ 2 (λ _ → AbGroupStr.is-set (snd H))) +-- λ _ → isOfHLevelPath 2 (AbGroupStr.is-set (snd H)) _ _ +-- isOfHLevel↑∙' zero (suc zero) = +-- subst (isOfHLevel 2) (sym (isoToPath (EM₁→∙Iso 0))) +-- (isSetΠ λ _ → emsquash _ _) +-- isOfHLevel↑∙' zero (suc (suc m)) = isOfHLevel↑∙-lem zero (suc m) +-- isOfHLevel↑∙' {H = H} (suc n) zero = +-- isOfHLevelΣ (2 + suc n) (isOfHLevelΠ (2 + suc n) +-- (λ _ → subst (isOfHLevel (suc (suc (suc n)))) +-- (cong (EM H) (cong suc (+-comm 0 n))) +-- (hLevelEM H (suc n)))) +-- λ _ → isOfHLevelPath (suc (suc (suc n))) +-- (subst (isOfHLevel (suc (suc (suc n)))) +-- (cong (EM H) (cong suc (+-comm 0 n))) +-- (hLevelEM H (suc n))) _ _ +-- isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc zero) = +-- subst (isOfHLevel (2 + suc n)) (sym (isoToPath (EM₁→∙Iso (suc n))) +-- ∙ λ i → EM-raw'∙ G 1 →∙ EM∙ H (suc (+-comm 1 n i))) +-- (isOfHLevelΠ (2 + suc n) λ x → (isOfHLevelTrunc (4 + n) _ _)) +-- isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc (suc m)) = +-- subst (isOfHLevel (2 + suc n)) +-- (λ i → (EM-raw'∙ G (suc (suc m)) +-- →∙ EM∙ H (suc (+-suc n (suc m) (~ i))))) +-- (isOfHLevel↑∙-lem (suc n) (suc m)) + +-- →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) +-- → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) +-- →∙EMPath {G = G} A n = +-- ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl +-- ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) + +-- private +-- contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) +-- → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) +-- fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl +-- snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = +-- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) +-- (sym (funExt (help n f p))) +-- where +-- help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) +-- → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) +-- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) +-- help zero f p = +-- raw-elim G zero +-- (λ _ → isOfHLevel↑∙ zero m _ _) p +-- help (suc n) f p = +-- Trunc.elim (λ _ → isOfHLevelPath (4 + n) +-- (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) +-- (λ i → suc (suc (+-comm (suc n) 1 i))) +-- (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) +-- (raw-elim G (suc n) +-- (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) + +-- contr∙-lem' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) +-- → isContr (EM∙ G (suc n) →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) +-- fst (contr∙-lem' n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +-- snd (contr∙-lem' {G = G} {H = H} {L = L} n m) (f , p) = +-- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) +-- (funExt λ x → sym (help' n f p x)) +-- where +-- help' : (n : ℕ) → (f : EM G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) +-- → f (snd (EM∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) +-- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) +-- help' zero f p = +-- raw-elim _ zero (λ _ → isOfHLevel↑∙' zero (suc m) _ _) p +-- help' (suc n) f p = +-- Trunc.elim (λ _ → isOfHLevelPath (4 + n) +-- (subst2 (λ x y → isOfHLevel x (EM-raw'∙ H (suc m) →∙ EM∙ L y)) +-- (λ i → suc (suc (suc (+-comm n 1 i)))) +-- (cong suc (+-suc n m)) +-- (isOfHLevelPlus' {n = 1} (suc (suc (suc n))) +-- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m)))) _ _) +-- (raw-elim _ (suc n) (λ _ → isOfHLevelPath' (2 + n) +-- (subst (λ y → isOfHLevel (suc (suc (suc n))) +-- (EM-raw'∙ H (suc m) →∙ EM∙ L y)) +-- (+-suc (suc n) m) +-- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p) + +-- contr∙-lem'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) +-- → isContr (EM-raw'∙ G (suc n) +-- →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) +-- fst (contr∙-lem'' n m) = (λ _ → (λ _ → 0ₖ (suc (n + m))) , refl) , refl +-- snd (contr∙-lem'' {G = G} {H = H} {L = L} n m) (f , p) = +-- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) +-- (funExt λ x → sym (help n f p x)) +-- where +-- help : (n : ℕ) → (f : EM-raw' G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) +-- → f (snd (EM-raw'∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) +-- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) +-- help zero f p = +-- EM-raw'-trivElim G zero (λ _ → isOfHLevel↑∙' _ _ _ _) p +-- help (suc n) f p = +-- EM-raw'-trivElim _ _ (λ _ → isOfHLevelPath' (suc (suc n)) +-- (subst (λ y → isOfHLevel (suc (suc (suc n))) (EM-raw'∙ H (suc m) →∙ EM∙ L y)) +-- (cong suc (+-suc n m)) +-- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p + +-- isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} +-- → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) +-- →∙ (EM∙ H (suc m) +-- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) +-- isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m zero = +-- isOfHLevelΩ→isOfHLevel 0 λ f +-- → subst +-- isProp (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) +-- where +-- h : isProp (EM∙ G (suc n) +-- →∙ (Ω (EM∙ H (suc m) +-- →∙ EM∙ L (suc (suc (n + m))) ∙))) +-- h = +-- subst isProp +-- (λ i → EM∙ G (suc n) +-- →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) +-- (isContr→isProp (contr∙-lem n m)) +-- isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = +-- isOfHLevelΩ→isOfHLevel (suc l) +-- λ f → +-- subst (isOfHLevel (2 + l)) +-- (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) +-- where +-- h : isOfHLevel (2 + l) +-- (EM∙ G (suc n) +-- →∙ (Ω (EM∙ H (suc m) +-- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) +-- h = +-- subst (isOfHLevel (2 + l)) +-- (λ i → EM∙ G (suc n) +-- →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) +-- (isOfHLevel↑∙∙ n m l) + +-- isOfHLevel↑∙∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} +-- → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) +-- →∙ (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) +-- isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m zero = +-- isOfHLevelΩ→isOfHLevel 0 +-- λ f → subst +-- isProp (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) +-- lem) +-- where +-- lem : isProp (EM∙ G (suc n) +-- →∙ (Ω (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (n + m))) ∙))) +-- lem = subst isProp +-- (λ i → EM∙ G (suc n) +-- →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) +-- (isContr→isProp (contr∙-lem' n m)) +-- isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m (suc l) = +-- isOfHLevelΩ→isOfHLevel (suc l) +-- λ f → +-- subst (isOfHLevel (2 + l)) +-- (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) +-- where +-- lem : isOfHLevel (2 + l) +-- (EM∙ G (suc n) +-- →∙ (Ω (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) +-- lem = +-- subst (isOfHLevel (2 + l)) +-- (λ i → EM∙ G (suc n) +-- →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) +-- (isOfHLevel↑∙∙' n m l) + +-- isOfHLevel↑∙∙'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} +-- → ∀ n m l → isOfHLevel (2 + l) (EM-raw'∙ G (suc n) +-- →∙ (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) +-- isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m zero = +-- isOfHLevelΩ→isOfHLevel 0 +-- λ f → subst +-- isProp (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) +-- lem) +-- where +-- lem : isProp (EM-raw'∙ G (suc n) +-- →∙ (Ω (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (n + m))) ∙))) +-- lem = subst isProp +-- (λ i → EM-raw'∙ G (suc n) +-- →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) +-- (isContr→isProp (contr∙-lem'' _ _)) +-- isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m (suc l) = +-- isOfHLevelΩ→isOfHLevel (suc l) +-- λ f → +-- subst (isOfHLevel (2 + l)) +-- (cong (λ x → typ (Ω x)) +-- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) +-- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) +-- where +-- lem : isOfHLevel (2 + l) +-- (EM-raw'∙ G (suc n) +-- →∙ (Ω (EM-raw'∙ H (suc m) +-- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) +-- lem = +-- subst (isOfHLevel (2 + l)) +-- (λ i → EM-raw'∙ G (suc n) +-- →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) +-- (isOfHLevel↑∙∙'' n m l) + +-- -- A homomorphism φ : G → H of AbGroups induces a homomorphism +-- -- φ' : K(G,n) → K(H,n) +-- inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- → AbGroupHom G' H' +-- → ∀ n +-- → EM-raw G' n → EM-raw H' n +-- inducedFun-EM-raw f = +-- elim+2 (fst f) +-- (EMrec _ emsquash embase +-- (λ g → emloop (fst f g)) +-- λ g h → compPathR→PathP (sym +-- (sym (lUnit _) +-- ∙∙ cong (_∙ (sym (emloop (fst f h)))) +-- (cong emloop (IsGroupHom.pres· (snd f) g h) +-- ∙ emloop-comp _ (fst f g) (fst f h)) +-- ∙∙ sym (∙assoc _ _ _) +-- ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) +-- ∙∙ sym (rUnit _)))) +-- (λ n ind → λ { north → north +-- ; south → south +-- ; (merid a i) → merid (ind a) i} ) + +-- inducedFun-EM-raw-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM-raw G' n) +-- → inducedFun-EM-raw (idGroupHom {G = AbGroup→Group G'}) n x ≡ x +-- inducedFun-EM-raw-id zero x = refl +-- inducedFun-EM-raw-id (suc zero) = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) +-- λ { embase-raw → refl ; (emloop-raw g i) → refl} +-- inducedFun-EM-raw-id (suc (suc n)) north = refl +-- inducedFun-EM-raw-id (suc (suc n)) south = refl +-- inducedFun-EM-raw-id (suc (suc n)) (merid a i) j = merid (inducedFun-EM-raw-id (suc n) a j) i + +-- inducedFun-EM-raw-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} +-- (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) +-- → (x : EM-raw G' n) → inducedFun-EM-raw (compGroupHom ϕ ψ) n x +-- ≡ inducedFun-EM-raw ψ n (inducedFun-EM-raw ϕ n x) +-- inducedFun-EM-raw-comp ϕ ψ zero x = refl +-- inducedFun-EM-raw-comp ϕ ψ (suc zero) = +-- EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) +-- λ { embase-raw → refl ; (emloop-raw g i) → refl} +-- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) north = refl +-- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) south = refl +-- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) (merid a i) j = +-- merid (inducedFun-EM-raw-comp ϕ ψ (suc n) a j) i + +-- inducedFun-EM : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- → AbGroupHom G' H' +-- → ∀ n +-- → EM G' n → EM H' n +-- inducedFun-EM f zero = inducedFun-EM-raw f zero +-- inducedFun-EM f (suc zero) = inducedFun-EM-raw f (suc zero) +-- inducedFun-EM f (suc (suc n)) = Trunc.map (inducedFun-EM-raw f (2 + n)) + +-- EM-raw→EM-funct : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- (n : ℕ) (ψ : AbGroupHom G H) (y : EM-raw G n) +-- → EM-raw→EM _ _ (inducedFun-EM-raw ψ n y) +-- ≡ inducedFun-EM ψ n (EM-raw→EM _ _ y) +-- EM-raw→EM-funct zero ψ y = refl +-- EM-raw→EM-funct (suc zero) ψ y = refl +-- EM-raw→EM-funct (suc (suc n)) ψ y = refl + +-- inducedFun-EM-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM G' n) +-- → inducedFun-EM (idGroupHom {G = AbGroup→Group G'}) n x ≡ x +-- inducedFun-EM-id zero x = refl +-- inducedFun-EM-id (suc zero) = inducedFun-EM-raw-id (suc zero) +-- inducedFun-EM-id (suc (suc n)) = +-- Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) +-- λ x → cong ∣_∣ₕ (inducedFun-EM-raw-id _ x) + +-- inducedFun-EM-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} +-- (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) +-- → (x : EM G' n) → inducedFun-EM (compGroupHom ϕ ψ) n x +-- ≡ inducedFun-EM ψ n (inducedFun-EM ϕ n x) +-- inducedFun-EM-comp ϕ ψ zero x = refl +-- inducedFun-EM-comp ϕ ψ (suc zero) = inducedFun-EM-raw-comp ϕ ψ (suc zero) +-- inducedFun-EM-comp ϕ ψ (suc (suc n)) = +-- Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) +-- λ x → cong ∣_∣ₕ (inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) x) + +-- inducedFun-EM0ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {ϕ : AbGroupHom G' H'} (n : ℕ) +-- → inducedFun-EM ϕ n (0ₖ n) ≡ 0ₖ n +-- inducedFun-EM0ₖ {ϕ = ϕ} zero = IsGroupHom.pres1 (snd ϕ) +-- inducedFun-EM0ₖ (suc zero) = refl +-- inducedFun-EM0ₖ (suc (suc n)) = refl + +-- inducedFun-EM-pres+ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- (ϕ : AbGroupHom G' H') (n : ℕ) (x y : EM G' n) +-- → inducedFun-EM ϕ n (x +ₖ y) ≡ inducedFun-EM ϕ n x +ₖ inducedFun-EM ϕ n y +-- inducedFun-EM-pres+ₖ ϕ zero x y = IsGroupHom.pres· (snd ϕ) x y +-- inducedFun-EM-pres+ₖ {G' = G'} {H' = H'} ϕ (suc n) = +-- EM-elim2 (suc n) (λ _ _ → isOfHLevelPath (2 + suc n) (hLevelEM _ (suc n)) _ _) +-- (wedgeConEM.fun _ _ n n +-- (λ _ _ → isOfHLevelPath' (suc n + suc n) +-- (subst (λ m → isOfHLevel (suc (suc m)) (EM H' (suc n))) +-- (sym (+-suc n n)) +-- (isOfHLevelPlus' {n = n} (3 + n) +-- (hLevelEM _ (suc n)))) +-- _ _) +-- (l n) +-- (r n) +-- (l≡r n)) +-- where +-- lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) → EM-raw→EM G (suc n) ptEM-raw ≡ 0ₖ _ +-- lem zero = refl +-- lem (suc n) = refl + +-- l : (n : ℕ) (y : EM-raw G' (suc n)) +-- → inducedFun-EM ϕ (suc n) ((EM-raw→EM G' (suc n) ptEM-raw) +-- +ₖ EM-raw→EM G' (suc n) y) +-- ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) +-- +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y) +-- l n y = (cong (inducedFun-EM ϕ (suc n)) +-- (cong (_+ₖ EM-raw→EM G' (suc n) y) (lem n) +-- ∙ lUnitₖ (suc n) (EM-raw→EM G' (suc n) y))) +-- ∙∙ sym (lUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) +-- ∙∙ cong (_+ₖ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) +-- (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) +-- ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) + +-- r : (n : ℕ) (x : EM-raw G' (suc n)) +-- → inducedFun-EM ϕ (suc n) +-- (EM-raw→EM G' (suc n) x +ₖ EM-raw→EM G' (suc n) ptEM-raw) +-- ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) +-- +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) +-- r n x = cong (inducedFun-EM ϕ (suc n)) +-- (cong (EM-raw→EM G' (suc n) x +ₖ_) (lem n) +-- ∙ rUnitₖ (suc n) (EM-raw→EM G' (suc n) x)) +-- ∙∙ sym (rUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x))) +-- ∙∙ cong (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) +ₖ_) +-- (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) +-- ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) + +-- l≡r : (n : ℕ) → l n ptEM-raw ≡ r n ptEM-raw +-- l≡r zero = refl +-- l≡r (suc n) = refl + +-- inducedFun-EM-pres-ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- (ϕ : AbGroupHom G' H') (n : ℕ) (x : EM G' n) +-- → inducedFun-EM ϕ n (-ₖ x) ≡ -ₖ (inducedFun-EM ϕ n x) +-- inducedFun-EM-pres-ₖ ϕ n = +-- morphLemmas.distrMinus _+ₖ_ _+ₖ_ +-- (inducedFun-EM ϕ n) (inducedFun-EM-pres+ₖ ϕ n) (0ₖ n) (0ₖ n) (-ₖ_) (-ₖ_) +-- (lUnitₖ n) (rUnitₖ n) +-- (lCancelₖ n) (rCancelₖ n) +-- (assocₖ n) (inducedFun-EM0ₖ n) + +-- EMFun-EM→ΩEM+1 : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- {ϕ : AbGroupHom G H} (n : ℕ) (x : EM G n) +-- → PathP (λ i → inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i) +-- ≡ inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i)) +-- (EM→ΩEM+1 n (inducedFun-EM ϕ n x)) +-- (cong (inducedFun-EM ϕ (suc n)) (EM→ΩEM+1 n x)) +-- EMFun-EM→ΩEM+1 {ϕ = ϕ} zero x = refl +-- EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc zero) x = +-- cong-∙ ∣_∣ₕ (merid (inducedFun-EM ϕ (suc zero) x)) +-- (sym (merid embase)) +-- ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc zero))) +-- (cong ∣_∣ₕ (merid x)) (cong ∣_∣ₕ (sym (merid embase)))) +-- ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc zero)))) +-- (sym (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase)))) +-- EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc (suc n)) = +-- Trunc.elim (λ _ → isOfHLevelPath (4 + n) +-- (isOfHLevelTrunc (5 + n) _ _) _ _) +-- λ a → cong-∙ ∣_∣ₕ (merid (inducedFun-EM-raw ϕ (2 + n) a)) +-- (sym (merid north)) +-- ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc (suc n)))) +-- (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (sym (merid north)))) +-- ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc (suc n))))) +-- (sym (cong-∙ ∣_∣ₕ (merid a) (sym (merid north)))) + + +-- inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- → AbGroupEquiv G' H' +-- → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +-- Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +-- Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +-- Iso.rightInv (inducedFun-EM-rawIso e n) = h n +-- where +-- h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) +-- (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) +-- h = elim+2 +-- (secEq (fst e)) +-- (elimSet _ (λ _ → emsquash _ _) refl +-- (λ g → compPathR→PathP +-- (sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) +-- (sym (lUnit _)) +-- ∙ rCancel _)))) +-- λ n p → λ { north → refl +-- ; south → refl +-- ; (merid a i) k → merid (p a k) i} +-- Iso.leftInv (inducedFun-EM-rawIso e n) = h n +-- where +-- h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) +-- (Iso.inv (inducedFun-EM-rawIso e n)) +-- h = elim+2 +-- (retEq (fst e)) +-- (elimSet _ (λ _ → emsquash _ _) refl +-- (λ g → compPathR→PathP +-- ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) +-- ∙ rCancel _))))) +-- λ n p → λ { north → refl +-- ; south → refl +-- ; (merid a i) k → merid (p a k) i} + +-- module _ {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) where +-- Iso→EMIso : ∀ n → Iso (EM G n) (EM H n) +-- Iso.fun (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom e) n +-- Iso.inv (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom (invGroupEquiv e)) n +-- Iso.rightInv (Iso→EMIso zero) = Iso.rightInv (inducedFun-EM-rawIso e zero) +-- Iso.rightInv (Iso→EMIso (suc zero)) = +-- Iso.rightInv (inducedFun-EM-rawIso e (suc zero)) +-- Iso.rightInv (Iso→EMIso (suc (suc n))) = +-- Iso.rightInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) +-- Iso.leftInv (Iso→EMIso zero) = +-- Iso.leftInv (inducedFun-EM-rawIso e zero) +-- Iso.leftInv (Iso→EMIso (suc zero)) = +-- Iso.leftInv (inducedFun-EM-rawIso e (suc zero)) +-- Iso.leftInv (Iso→EMIso (suc (suc n))) = +-- Iso.leftInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) + +-- Iso→EMIso∙ : ∀ n → Iso.fun (Iso→EMIso n) (EM∙ G n .snd) ≡ EM∙ H n .snd +-- Iso→EMIso∙ zero = IsGroupHom.pres1 (e .snd) +-- Iso→EMIso∙ (suc zero) = refl +-- Iso→EMIso∙ (suc (suc n)) = refl + +-- Iso→EMIso⁻∙ : ∀ n → Iso.inv (Iso→EMIso n) (EM∙ H n .snd) ≡ EM∙ G n .snd +-- Iso→EMIso⁻∙ zero = IsGroupHom.pres1 (invGroupEquiv e .snd) +-- Iso→EMIso⁻∙ (suc zero) = refl +-- Iso→EMIso⁻∙ (suc (suc n)) = refl + +-- Iso→EMIsoInv : {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) +-- → ∀ n → Iso.inv (Iso→EMIso e n) ≡ Iso.fun (Iso→EMIso (invGroupEquiv e) n) +-- Iso→EMIsoInv e zero = refl +-- Iso→EMIsoInv e (suc zero) = refl +-- Iso→EMIsoInv e (suc (suc n)) = refl + +-- EM⊗-commIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} +-- → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +-- EM⊗-commIso {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +-- EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} +-- → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +-- EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 6c1fa607b7..fdf4dd9f7b 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -51,7 +51,7 @@ module example (G G* G○ : Group ℓ) lhs rhs : fst G - lhs = (p · p') · (inv p' · (F* ⟪ (((*.inv q) *.· r) *.· F○ ⟪ ○.inv t ⟫ *.· + lhs = p · (p · inv p) · inv p · (p' · inv p') · (p · p') · (inv p' · (F* ⟪ (((*.inv q) *.· r) *.· F○ ⟪ ○.inv t ⟫ *.· (*.inv (F○ ⟪ s ○.· s ⟫) *.· F○ ⟪ u ⟫ )) ⟫ )) rhs = inv (F* ⟪ q ⟫ · inv p) · (F* ⟪ r ⟫) · diff --git a/Cubical/Tactics/PathSolver/ReflectionTest.agda b/Cubical/Tactics/PathSolver/ReflectionTest.agda new file mode 100644 index 0000000000..b32b5c44b0 --- /dev/null +++ b/Cubical/Tactics/PathSolver/ReflectionTest.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.ReflectionTest where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + + +module Test1 where + + module _ (n : ℕ) where + data A : Type where + a aa : A + p : a ≡ aa + + macro + unquoteA : R.Term → R.TC Unit + unquoteA hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + final ← R.checkType (R.def (quote refl) []) ty >>= R.normalise + R.unify hole final + -- (R∙ (R.def (quote refl) []) (R.def (quote refl) [])) + + a' : Path (Path (A 3) a aa) (refl ∙ p) (refl ∙ p) + a' = unquoteA diff --git a/Cubical/Tactics/PathSolver/Solver2.agda b/Cubical/Tactics/PathSolver/Solver2.agda index fc8e7f0abf..965d207851 100644 --- a/Cubical/Tactics/PathSolver/Solver2.agda +++ b/Cubical/Tactics/PathSolver/Solver2.agda @@ -636,12 +636,12 @@ macro -- solveSquare = squareSolverMain false -module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where +-- module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - pA pB pC : x ≡ w - pA = (p ∙∙ refl ∙∙ q) ∙ sym r - pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r +-- pA pB pC : x ≡ w +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym r +-- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl +-- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - pA=pB : pA ≡ pB - pA=pB = solveGroupoidDebug +-- pA=pB : pA ≡ pB +-- pA=pB = solveGroupoidDebug diff --git a/Cubical/Tactics/PathSolver/Solver3.agda b/Cubical/Tactics/PathSolver/Solver3.agda new file mode 100644 index 0000000000..f213d278ea --- /dev/null +++ b/Cubical/Tactics/PathSolver/Solver3.agda @@ -0,0 +1,806 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Solver3 where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + +private + variable + ℓ ℓ' : Level + A A' : Type ℓ + + +infixr 5 _∷Tω_ + +data [Typeₙ] : Typeω where + [] : [Typeₙ] + _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] + +maxℓ : [Typeₙ] → Level +maxℓ [] = ℓ-zero +maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) + +lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) +lookupTₙ [] x = ⊥* +lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ +lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) + +ℓAt : (Ts : [Typeₙ]) → ℕ → Level +ℓAt [] x = ℓ-zero +ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ +ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x + +TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) +TyAt [] k = ⊥* +TyAt (x ∷Tω Ts) zero = x +TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k + +cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k +cast↓ [] k () +cast↓ (x₁ ∷Tω Ts) zero x = lower x +cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) + +cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y +cast↓Inj {[]} {A = A} {()} +cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift +cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = + cong lift (cast↓Inj {[T] = [T]} {A = A} p) + +cast↓Inj-sec : ∀ {[T] A x y} p → + cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p +cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl +cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = + cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) + +cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → + cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) + ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) + + +cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ +cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = + cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ + + +cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k +cast↑ [] k () +cast↑ (x₁ ∷Tω Ts) zero x = lift x +cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x + +-- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) +-- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t + + +-- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) +-- Ts-arr' [] x x₁ = Unit +-- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) +-- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} +-- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} +-- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = +-- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) + +-- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → +-- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t +-- Ts-arr' Ts s t x x₁ = {!!} + + + +data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) + compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) + congCase : ∀ {ℓ ℓ' A A'} (f : A → A') {x y} + → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) + + +module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} + (p : Path {ℓ} A x y) where + + pathCaseCongTest : PathCase (cong f p) + pathCaseCongTest = congCase f λ z → p z + + +infixl 5 _fp∷_ +infixl 5 _fp++_ + +data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where + [] : ∀ {x b} → FlatPath A b x x + _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z + _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y + + +FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ +FP⟦ [] ⟧ = refl +FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ + +pattern fp[_] x = [] fp∷ x + +_fp++_ : ∀ {x y z} + → FlatPath A false x y + → FlatPath A false y z + → FlatPath A false x z +x fp++ [] = x +x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ + +fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) + (fp' : FlatPath A false a₁ a₂) + → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ +fp++∙ fp [] = sym (rUnit _) +fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') + +module _ ([T] : [Typeₙ]) where + + data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where + reflExpr : ∀ {A x} → PathExpr A x x + atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y + compExpr : ∀ {A x y z w} + → PathExpr A x y → PathExpr A y z → PathExpr A z w + → PathExpr A x w + congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') + → PathExpr A x y + → PathExpr A' (f x) (f y) + + + PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → + (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) + PE⟦ reflExpr ⟧ = refl + PE⟦ atomExpr p ⟧ = cong _ p + PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ + PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) + + cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) + → PathExpr A₁ a₀ a₁ + + → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) + (cast↓ [T] A (f a₁)) + cong-flat f reflExpr = [] + cong-flat f (atomExpr p) = fp[ cong _ p ] + cong-flat f (compExpr x x₁ x₂) = + cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ + cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x + + flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ + → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) + flat⟦ reflExpr ⟧ = [] + flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] + flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ + flat⟦ congExpr f x ⟧ = cong-flat f x + + + cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → + (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → + (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ + FP⟦ cong-flat f pe ⟧ + cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) + cong-flat≡ (atomExpr p) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ + cong-flat≡ (compExpr pe pe₁ pe₂) f = + (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ + cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ + (λ i → doubleCompPath-elim + (cong-flat≡ pe f i) + (cong-flat≡ pe₁ f i) + (cong-flat≡ pe₂ f i) i)) + ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) + (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) + ∙∙ fp++∙ _ (cong-flat f pe₂) + cong-flat≡ (congExpr f₁ pe) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) + + pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → + PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ + pe≡flat reflExpr = refl + pe≡flat (atomExpr p) = lUnit _ + pe≡flat (compExpr pe pe₁ pe₂) = + (λ i → doubleCompPath-elim + (pe≡flat pe i) + (pe≡flat pe₁ i) + (pe≡flat pe₂ i) i) + ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) + ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ + + pe≡flat (congExpr f pe) = cong-flat≡ pe f + +-- data PathExprω : {ℓ : Level} (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where +-- reflExpr : ∀ {ℓ A x} → PathExprω {ℓ} A (refl {x = x}) +-- atomExpr : ∀ {ℓ A x y} → (p : x ≡ y) → PathExprω {ℓ} A p +-- compExpr : ∀ {ℓ A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} +-- → PathExprω {ℓ} A p → PathExprω {ℓ} A q → PathExprω {ℓ} A r +-- → PathExprω {ℓ} A (p ∙∙ q ∙∙ r) +-- congExpr : ∀ {ℓ ℓ' A A' x y} → (f : A → A') {p : x ≡ y} +-- → PathExprω {ℓ} A p +-- → PathExprω {ℓ'} A' (cong f p) + + +-- data PathExpr {ℓ : Level} : (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Type (ℓ-suc ℓ) where +-- reflExpr : ∀ {A x} → PathExpr A (refl {x = x}) +-- atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A p +-- compExpr : ∀ {A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} +-- → PathExpr A p → PathExpr A q → PathExpr A r +-- → PathExpr A (p ∙∙ q ∙∙ r) +-- congExpr : ∀ {A A' x y} → (f : A → A') {p : x ≡ y} +-- → PathExpr {ℓ} A {x} {y} p +-- → PathExpr A' (cong f p) + + + +-- PEω⟦_⟧ : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → +-- PathExprω A p → a₀ ≡ a₁ +-- PEω⟦ reflExpr ⟧ = refl +-- PEω⟦ atomExpr p ⟧ = p +-- PEω⟦ compExpr x x₁ x₂ ⟧ = PEω⟦ x ⟧ ∙∙ PEω⟦ x₁ ⟧ ∙∙ PEω⟦ x₂ ⟧ +-- PEω⟦ congExpr f x ⟧ = cong f PEω⟦ x ⟧ + + + +-- record ΣPathExprω : Typeω where +-- field +-- ℓa : Level +-- A : Type ℓa +-- a₀ a₁ : A +-- p : a₀ ≡ a₁ +-- expr : PathExprω A p + +-- ΣPEω⟦_⟧ : a₀ ≡ a₁ +-- ΣPEω⟦_⟧ = PEω⟦ expr ⟧ + +-- record ΣPathExpr {ℓ} : Type (ℓ-suc ℓ) where +-- constructor Σpe +-- field +-- {A} : Type ℓ +-- {a₀ a₁} : A +-- {p} : a₀ ≡ a₁ +-- expr : PathExpr A p + +-- ΣPE⟦_⟧ : a₀ ≡ a₁ +-- ΣPE⟦_⟧ = PE⟦ expr ⟧ + + + +module PathTrm (A B : Type ℓ) where + data PathTrm : Type ℓ where + reflTrm : PathTrm + atomTrm : A → PathTrm + compTrm : PathTrm → PathTrm → PathTrm → PathTrm + congTrm : B → PathTrm → PathTrm + + module showPathTrm (showA : A → String) (showB : B → String) where + showPT : PathTrm → String + showPT reflTrm = "refl" + showPT (atomTrm x) = showA x + showPT (compTrm x x₁ x₂) = + "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" + showPT (congTrm x x₁) = + "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" + + +module _ {ℓ ℓ'} + {A B : Type ℓ} + {A' B' : Type ℓ'} + (f : A → R.TC A') + (g : B → R.TC B') where + open PathTrm + mmapPT : PathTrm A B → R.TC $ PathTrm A' B' + mmapPT reflTrm = pure reflTrm + mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ + mmapPT (compTrm x x₁ x₂) = + ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ + mmapPT (congTrm x x₁) = + ⦇ congTrm (g x) (mmapPT x₁) ⦈ + + +module RTrm = PathTrm R.Term R.Term +module StrTrm = PathTrm String String + +-- unqouteRTrm : ∀ {ℓ} → RTrm.PathTrm → R.TC (ΣPathExpr {ℓ}) +-- unqouteRTrm PathTrm.reflTrm = pure $ Σpe {A = Unit*} reflExpr +-- unqouteRTrm (PathTrm.atomTrm x) = +-- ⦇ (Σpe ∘ atomExpr) {!x!} ⦈ +-- unqouteRTrm (PathTrm.compTrm x x₁ x₂) = {!!} +-- unqouteRTrm (PathTrm.congTrm x x₁) = {!!} + +showRTrmTC : RTrm.PathTrm → R.TC String +showRTrmTC = + mmapPT + (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + + + +module tryPathE where + + try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm) + + + tryRefl : R.Term → R.TC (RTrm.PathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + R.returnTC (RTrm.reflTrm) + + tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (t1') ← try≡ k t1 + (t2') ← try≡ k t2 + (t3') ← try≡ k t3 + R.returnTC (RTrm.compTrm t1' t2' t3') + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + atom : R.Term → R.TC (RTrm.PathTrm) + atom x = R.returnTC (RTrm.atomTrm x) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) (atom t)) + + + + + + + + + + -- data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + -- isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) + -- isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p + -- isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} + -- → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) + + -- data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where + -- nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) + -- consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) + + -- data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where + -- nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x + -- consTrmInvol : ∀ {x y z : _} → + -- IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z + -- involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → + -- IsPathTrmInvol x y + + +-- module show (showA : ∀ {ℓ} A {a₀ a₁ : A} → (p : a₀ ≡ a₁) → String) where +-- showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String +-- showIPT isReflTrm = "refl" +-- showIPT (isAtomTrm x) = showA x +-- showIPT (isCompTrm isReflTrm x₁ x₂) = +-- "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" +-- showIPT (isCompTrm x x₁ isReflTrm) = +-- "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" +-- showIPT (isCompTrm x x₁ x₂) = +-- "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" + +-- -- showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String + +-- -- showIPTD nilTrmDeasRefl = "refl" +-- -- showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q + +-- -- showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String +-- -- showIPTI nilTrmInvolRefl = "refl" +-- -- showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q +-- -- showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" + + +-- -- depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} +-- -- → IsPathTrmDeas p → ℕ +-- -- depthIsPathTrmDeas nilTrmDeasRefl = zero +-- -- depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) + +-- -- hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool +-- -- hasRedexes nilTrmInvolRefl = false +-- -- hasRedexes (consTrmInvol x q) = hasRedexes x +-- -- hasRedexes (involTrmInvol x q) = true + +-- -- Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ +-- -- Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl +-- -- Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q + +-- -- IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → +-- -- IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q +-- -- IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' +-- -- IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' +-- -- IsPathTrmDeas∙ p' (consTrmDeas q' q'') = +-- -- let (_ , u) = IsPathTrmDeas∙ p' q' +-- -- in _ , consTrmDeas u q'' + + +-- -- Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} + +-- -- Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl +-- -- Invol→Deas↓ (consTrmInvol x q) = +-- -- IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) +-- -- Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x + +-- -- ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) +-- -- ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl +-- -- ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl +-- -- ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = +-- -- let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r +-- -- in _ , isCompTrmReg (isCompTrmReg pR qR) rR , +-- -- λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i + + +-- -- ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) +-- -- ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl +-- -- ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p +-- -- ⟦ isCompTrmReg p' q' ⟧da = +-- -- let (_ , qD) = ⟦ q' ⟧da +-- -- (_ , pD) = ⟦ p' ⟧da +-- -- (_ , p∙qD) = IsPathTrmDeas∙ pD qD +-- -- in _ , p∙qD + +-- -- ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) +-- -- ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da +-- -- module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) +-- -- (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where + +-- -- IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → +-- -- (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → +-- -- (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) +-- -- IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ +-- -- IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ +-- -- IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = +-- -- cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ +-- -- sym (∼assoc _ _ _) + +-- -- ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → +-- -- fst ⟦ p' ⟧da ≡ p +-- -- ⟦ isReflTrmReg ⟧da≡ = refl +-- -- ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) +-- -- ⟦ isCompTrmReg p' q' ⟧da≡ = +-- -- IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ + +-- -- daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da +-- -- daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) + + + +-- -- module _ {A : Type ℓ} where + +-- -- open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public +-- -- open PT≡ rUnit lUnit public + + + +-- -- ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ +-- -- ⟦ nilTrmInvolRefl , _ ⟧ti = refl +-- -- ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q +-- -- ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q +-- -- ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti +-- -- ⟦ involTrmInvol x q , seg j ⟧ti i = +-- -- hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i +-- -- ; (i = i1) → q (~ k ∧ ~ j) +-- -- ; (i = i0) → ⟦ x , seg j ⟧ti i0 +-- -- }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) + +-- -- ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti +-- -- ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti + + + +-- module _ (A : Type ℓ) (a : A) where +-- module PTG = PT {A = Unit} (λ _ _ → A) +-- (a) +-- (λ _ _ → a) +-- (λ _ _ _ → a) +-- (λ _ _ _ → refl) +-- (λ _ _ _ → refl) + +-- module PTrm = PTG R.Term R.unknown + +-- module Pℕ = PTG (Bool × ℕ) (true , 0) + +-- module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) + + +-- -- module _ (f : (Bool × ℕ) → R.Term) where +-- -- mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ +-- -- mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl +-- -- mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) +-- -- mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) + + +-- IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) +-- IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ + +-- -- ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ + +-- -- consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ +-- -- consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x +-- -- consInvolℕ q'@(PT.consTrmDeas x q) x₁ = +-- -- decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) +-- -- (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) + + + +-- -- ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl +-- -- ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' + +-- module tryPathE where + +-- try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) + + +-- tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- tryRefl t = do +-- _ ← R.noConstraints $ R.checkType +-- (R.con (quote reflCase) []) +-- (R.def (quote PathCase) ([ varg t ])) +-- R.returnTC (_ , PTrm.isReflTrm) + +-- tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] +-- tryComp (suc k) t = do +-- tm ← R.noConstraints $ R.checkType +-- (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) +-- (R.def (quote PathCase) ([ varg t ])) +-- (t1 , t2 , t3) ← h tm +-- (_ , t1') ← try≡ k t1 +-- (_ , t2') ← try≡ k t2 +-- (_ , t3') ← try≡ k t3 +-- R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') + +-- where + +-- h : R.Term → R.TC (R.Term × R.Term × R.Term) +-- h (R.con (quote compCase) l) = match3Vargs l +-- h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + +-- atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) +-- atom x = R.returnTC (_ , PTrm.isAtomTrm x) + + +-- try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] +-- try≡ (suc k) t = +-- R.catchTC +-- (tryRefl t) +-- (R.catchTC (tryComp k t) (atom t)) + +-- lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y +-- lamWrap' p i = p i + + +-- lamWrap : R.Term → R.Term +-- lamWrap t = R.def (quote lamWrap') (t v∷ []) + +-- unLam : R.Term → R.TC R.Term +-- unLam (R.lam _ (R.abs _ t)) = R.returnTC t +-- unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) + +-- appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) +-- appendIfUniqe x [] = R.returnTC [ x ] +-- appendIfUniqe x xs@(x₁ ∷ xs') = do +-- x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) +-- x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) +-- ( +-- R.catchTC +-- (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) +-- (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) +-- ) + +-- comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) +-- comparePathTerms x x₁ = do +-- x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) +-- x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= +-- λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) +-- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam +-- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) +-- ( +-- R.catchTC +-- (R.extendContext "i" (varg (R.def (quote I) [])) $ +-- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) +-- (R.returnTC nothing) +-- ) + +-- concatUniq : List R.Term → List R.Term → R.TC (List R.Term) +-- concatUniq [] = R.returnTC +-- concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ + +-- indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) +-- indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) +-- indexOfAlways t xs@(x ∷ xs') = +-- comparePathTerms t x >>= +-- Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) +-- (λ b → R.returnTC (xs , (b , 0))) + +-- unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → +-- (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) +-- unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) +-- unMapAtoms l (PT.isAtomTrm x) = +-- do (l' , y) ← indexOfAlways x l +-- R.returnTC (l' , _ , Pℕ.isAtomTrm y) +-- unMapAtoms l (PT.isCompTrm e e₁ e₂) = +-- do (l' , _ , e') ← unMapAtoms l e +-- (l'' , _ , e₁') ← unMapAtoms l' e₁ +-- (l''' , _ , e₂') ← unMapAtoms l'' e₂ +-- (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) + + +-- -- unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term +-- -- unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] +-- -- unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) +-- -- unquotePathTrm (PT.isCompTrm x x₁ x₂) = +-- -- let x' = unquotePathTrm x +-- -- x₁' = unquotePathTrm x₁ +-- -- x₂' = unquotePathTrm x₂ +-- -- in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) + +-- -- module _ (l : List R.Term) where +-- -- lk : (Bool × ℕ) → R.Term +-- -- lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) +-- -- where +-- -- z = lookupWithDefault R.unknown l n + + + +-- -- mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) +-- -- mkTrmsInvol' zero _ = [] +-- -- mkTrmsInvol' (suc k) u = +-- -- let pi = ℕDeas→ℕInvol u +-- -- in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] + +-- -- mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) +-- -- mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x + +-- -- unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term +-- -- -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) +-- -- -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) +-- -- unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] +-- -- unquoteTrmInvol (PT.consTrmInvol x q) = +-- -- R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) +-- -- unquoteTrmInvol (PT.involTrmInvol x q) = +-- -- R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) + +-- -- mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) +-- -- mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) +-- -- ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x + +-- -- foldPathCompTerm : List R.Term → R.Term +-- -- foldPathCompTerm [] = R.def (quote refl) [] +-- -- foldPathCompTerm (x ∷ []) = x +-- -- foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) + +-- -- mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) +-- -- mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) + + +groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag hole = do + + (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + r0 ← tryPathE.try≡ 100 t0 + r1 ← tryPathE.try≡ 100 t1 + + + -- (aL' , (_ , e0)) ← unMapAtoms [] r0' + -- (aL , (_ , e1)) ← unMapAtoms aL' r1' + -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + show0 ← showRTrmTC r0 + show1 ← showRTrmTC r1 + + let dbgInfo = ("LHS: ") ∷ₑ show0 ∷nl + ("RHS: ") ∷ₑ show1 ∷nl [] + -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) + + -- ∷ (R.strErr "\n") + -- (niceAtomList aL) + + R.typeError dbgInfo + + +groupoidSolverMain : Bool → R.Term → R.TC Unit +groupoidSolverMain debugFlag hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + +squareSolverMain : Bool → R.Term → R.TC Unit +squareSolverMain debugFlag hole = do + ty ← R.inferType hole >>= wait-for-type + hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta + + (solution , msg) ← groupoidSolverTerm debugFlag hole' + R.catchTC + (R.unify hole' solution) + (R.typeError msg) + + R.catchTC + (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) + (R.typeError [ R.strErr "xxx" ] ) + where + extractMeta : R.Term → R.TC R.Term + extractMeta (R.def _ tl) = matchVarg tl + extractMeta t = + R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) + +macro + solveGroupoidDebug : R.Term → R.TC Unit + solveGroupoidDebug = groupoidSolverMain true + + -- solveSquareDebug : R.Term → R.TC Unit + -- solveSquareDebug = squareSolverMain false + + -- solveGroupoid : R.Term → R.TC Unit + -- solveGroupoid = groupoidSolverMain true + + -- solveSquare : R.Term → R.TC Unit + -- solveSquare = squareSolverMain false + + +-- module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where + +-- pA pB pC : x ≡ w +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym r +-- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl +-- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + +-- pA=pB : pA ≡ pB +-- pA=pB = solveGroupoidDebug diff --git a/Cubical/Tactics/PathSolver/Solver4.agda b/Cubical/Tactics/PathSolver/Solver4.agda new file mode 100644 index 0000000000..b5ad972dde --- /dev/null +++ b/Cubical/Tactics/PathSolver/Solver4.agda @@ -0,0 +1,877 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Solver4 where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + +private + variable + ℓ ℓ' : Level + A A' : Type ℓ + + +infixr 5 _∷Tω_ + +infixr 5 _∷fn_ + + +data [Typeₙ] : Typeω where + [Tω] : [Typeₙ] + _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] + +data [Fns] : Typeω where + [fn] : [Fns] + _∷fn_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] + + +reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) +reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] _ = + R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ + + +-- macro +-- test-refl[T]macro : R.Term → R.Term → R.TC Unit +-- test-refl[T]macro inp hole = do +-- l ← reflected[Typeₙ]→[reflectedTy] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[T] : Unit +-- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) + + +reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) +reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ + +-- macro +-- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit +-- test-refl[Fn]macro inp hole = do +-- l ← reflected[Fns]→[reflectedFn] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[Fn] : Unit +-- test-refl[Fn] = test-refl[Fn]macro +-- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) + + +maxℓ : [Typeₙ] → Level +maxℓ [Tω] = ℓ-zero +maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) + +lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) +lookupTₙ [Tω] x = ⊥* +lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ +lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) + +ℓAt : (Ts : [Typeₙ]) → ℕ → Level +ℓAt [Tω] x = ℓ-zero +ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ +ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x + +TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) +TyAt [Tω] k = ⊥* +TyAt (x ∷Tω Ts) zero = x +TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k + +cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k +cast↓ [Tω] k () +cast↓ (x₁ ∷Tω Ts) zero x = lower x +cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) + +cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y +cast↓Inj {[Tω]} {A = A} {()} +cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift +cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = + cong lift (cast↓Inj {[T] = [T]} {A = A} p) + +cast↓Inj-sec : ∀ {[T] A x y} p → + cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p +cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl +cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = + cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) + +cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → + cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) + ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) + + +cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ +cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = + cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ + + +cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k +cast↑ [Tω] k () +cast↑ (x₁ ∷Tω Ts) zero x = lift x +cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x + +-- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) +-- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t + + +-- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) +-- Ts-arr' [] x x₁ = Unit +-- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) +-- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} +-- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} +-- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = +-- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) + +-- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → +-- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t +-- Ts-arr' Ts s t x x₁ = {!!} + + + +data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) + compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) + congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') + → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) + + +module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} + (p : Path {ℓ} A x y) where + + -- pathCaseCongTest : PathCase λ i → f (p i) + -- pathCaseCongTest = congCase f {!!} + + +infixl 5 _fp∷_ +infixl 5 _fp++_ + +data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where + [fp] : ∀ {x b} → FlatPath A b x x + _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z + _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y + +magicNumber = 100 + +mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) +mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing +mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w + where + w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) + w (R.con (quote [fp]) _ , _) = R.returnTC nothing + w (xs'@(R.con (quote _fp∷_) tl') , y) = + match2Vargs tl' >>= λ (xs , x) → + R.catchTC + (R.noConstraints $ R.unify + (R.def (quote sym) (x v∷ [])) y + >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) + >>= λ (xs* , xs*↑) → + R.returnTC + (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) + , xs*↑ ) + )) + (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) + ((xs* v∷ y v∷ [])))) + <$> mb-invol n xs') + w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) +mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) + +mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol' = mb-invol magicNumber + + +redList : ℕ → R.Term → R.TC (List R.Term) +redList = h + where + h : ℕ → R.Term → R.TC (List R.Term) + h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) + h (suc k) t = + (mb-invol' t) >>= + Mb.rec + (R.returnTC []) + λ (t' , t'↓) → do + p' ← h k t'↓ + R.returnTC (t' ∷ p') + + +redList' : R.Term → R.TC (List R.Term) +redList' = redList magicNumber + + +pattern fp[_] x = [fp] fp∷ x + +FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ +FP⟦ [fp] ⟧ = refl +FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ + +-- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ +-- FP⟦ [] ⟧t = refl +-- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ +-- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ + + +fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ +fpF→T [fp] = [fp] +fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ + +fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ +fpT→F _ [fp] = [fp] +fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ +fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ +fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ + +fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → + FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ +fpT≡F [fp] = refl +fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x +fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = + hcomp + (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ k ∧ ~ i) + }) + (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ i ∧ k) + }) + (fpT≡F fp i j)) + +_fp++_ : ∀ {x y z} + → FlatPath A false x y + → FlatPath A false y z + → FlatPath A false x z +x fp++ [fp] = x +x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ + +fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) + (fp' : FlatPath A false a₁ a₂) + → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ +fp++∙ fp [fp] = sym (rUnit _) +fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') + +module PE ([T] : [Typeₙ]) where + + data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where + reflExpr : ∀ {A x} → PathExpr A x x + atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y + compExpr : ∀ {A x y z w} + → PathExpr A x y → PathExpr A y z → PathExpr A z w + → PathExpr A x w + congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') + → PathExpr A x y + → PathExpr A' (f x) (f y) + + + PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → + (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) + PE⟦ reflExpr ⟧ = refl + PE⟦ atomExpr p ⟧ = cong _ p + PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ + PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) + + cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) + → PathExpr A₁ a₀ a₁ + + → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) + (cast↓ [T] A (f a₁)) + cong-flat f reflExpr = [fp] + cong-flat f (atomExpr p) = fp[ cong _ p ] + cong-flat f (compExpr x x₁ x₂) = + cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ + cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x + + flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ + → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) + flat⟦ reflExpr ⟧ = [fp] + flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] + flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ + flat⟦ congExpr f x ⟧ = cong-flat f x + + + cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → + (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → + (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ + FP⟦ cong-flat f pe ⟧ + cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) + cong-flat≡ (atomExpr p) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ + cong-flat≡ (compExpr pe pe₁ pe₂) f = + (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ + cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ + (λ i → doubleCompPath-elim + (cong-flat≡ pe f i) + (cong-flat≡ pe₁ f i) + (cong-flat≡ pe₂ f i) i)) + ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) + (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) + ∙∙ fp++∙ _ (cong-flat f pe₂) + cong-flat≡ (congExpr f₁ pe) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) + + pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → + PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ + pe≡flat reflExpr = refl + pe≡flat (atomExpr p) = lUnit _ + pe≡flat (compExpr pe pe₁ pe₂) = + (λ i → doubleCompPath-elim + (pe≡flat pe i) + (pe≡flat pe₁ i) + (pe≡flat pe₂ i) i) + ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) + ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ + + pe≡flat (congExpr f pe) = cong-flat≡ pe f + + +module PathTrm (A B : Type ℓ) where + data PathTrm : Type ℓ where + reflTrm : PathTrm + atomTrm : A → PathTrm + compTrm : PathTrm → PathTrm → PathTrm → PathTrm + congTrm : B → PathTrm → PathTrm + + module showPathTrm (showA : A → String) (showB : B → String) where + showPT : PathTrm → String + showPT reflTrm = "refl" + showPT (atomTrm x) = showA x + showPT (compTrm x x₁ x₂) = + "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" + showPT (congTrm x x₁) = + "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" + + +module _ {ℓ ℓ'} + {A B : Type ℓ} + {A' B' : Type ℓ'} + (f : A → R.TC A') + (g : B → R.TC B') where + open PathTrm + mmapPT : PathTrm A B → R.TC $ PathTrm A' B' + mmapPT reflTrm = pure reflTrm + mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ + mmapPT (compTrm x x₁ x₂) = + ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ + mmapPT (congTrm x x₁) = + ⦇ congTrm (g x) (mmapPT x₁) ⦈ + + +module RTrm = PathTrm R.Term R.Term +module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) +module StrTrm = PathTrm String String + +showRTrmTC : RTrm.PathTrm → R.TC String +showRTrmTC = + mmapPT + (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +showRTrmTC' : RTrm'.PathTrm → R.TC String +showRTrmTC' = + let q = λ (k , t) → + R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) + in mmapPT + q q + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +module _ ([T] : [Typeₙ]) where + reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) + reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} + + + atomExpr' : ∀ A {x y} → (p : x ≡ y) → + PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) + + compExpr' : ∀ A {x y z w} + → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w + → PE.PathExpr [T] A x w + compExpr' A = PE.compExpr {[T] = [T]} {A = A} + + congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') + → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) + (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) + congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} + (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x + +getEnd : ∀ {x y : A} → x ≡ y → A +getEnd p = p i0 + +module tryPathE --([T] : [Typeₙ]) + (TC[T]trm : R.Term) + ([TC[T]trm] [fns] : List R.Term) where + + inferA : R.Term → R.TC ℕ + inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + + inferA' : R.Term → R.TC R.Term + inferA' = inferA >=> R.quoteTC + + + + try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + + + tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + let x₀ = R.def (quote getEnd) v[ t ] + A ← inferA' t + ⦇ (RTrm.reflTrm , R.def (quote reflExpr') + (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ + + tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (t1 , t1') ← (try≡ k t1) + (t2 , t2') ← (try≡ k t2) + (t3 , t3') ← (try≡ k t3) + A ← inferA' t + pure (RTrm.compTrm t1 t2 t3 , + R.def (quote compExpr') + (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t*) ← h tm + A ← inferA' t* + A' ← inferA' t + (t0 , t0') ← (try≡ k t*) + pure (RTrm.congTrm f t0 , + R.def (quote congExpr') + (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) + + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + + atom : R.Term → R.TC (RTrm.PathTrm × R.Term) + atom x = do + A ← inferA' x + pure (RTrm.atomTrm x , R.def (quote atomExpr') + (TC[T]trm v∷ A v∷ x v∷ [])) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) + + + +module tryTermE --([T] : [Typeₙ]) + --(TC[T]trm : R.Term) + ([TC[T]trm] [fns] : List R.Term) where + + try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) + + + tryRefl : R.Term → R.TC (RTrm'.PathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + ⦇ RTrm'.reflTrm ⦈ + + tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t) ← h tm + ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) + + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + + atom : R.Term → R.TC (RTrm'.PathTrm) + atom x = do + k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + R.returnTC (RTrm'.atomTrm (k , x)) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) + + + +-- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) +-- funTypeView = {!!} + +groupoidSolverTerm : Bool → R.Term → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag inp₀ inpF hole = do + + -- inp₀ ← wait-for-type inp₀' + + -- R.typeError (niceAtomList [) + (A' , (t0' , t1')) ← R.inferType hole >>= wait-for-type >>= (get-boundaryWithDom ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + -- A ← wait-for-type A' + -- t0 ← wait-for-type t0' + -- t1 ← wait-for-type t1' + + let t0 = t0' + let t1 = t1' + + (AA , _) ← get-boundaryWithDom A' >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary2" ]) + (λ x → R.returnTC x) + + let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) + -- (R.typeError [ inp ]ₑ) + + + + [t] ← reflected[Typeₙ]→[reflectedTy] inp + [f] ← reflected[Fns]→[reflectedFn] inpF + + (r0 , r0') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t0) + (r1 , r1') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t1) + + + + -- (aL' , (_ , e0)) ← unMapAtoms [] r0' + -- (aL , (_ , e1)) ← unMapAtoms aL' r1' + -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + show0 ← showRTrmTC r0 + show1 ← showRTrmTC r1 + + red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' + red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' + + + let invPa0 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red0 + invPa1 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red1 + + let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl + ("LHS: ") ∷ₑ show0 ∷nl + ("RHS: ") ∷ₑ show1 ∷nl + (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) + -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) + + -- ∷ (R.strErr "\n") + -- (niceAtomList aL) + + let finalTrm0 = + just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) + ∷ invPa0 + + finalTrm1 = + just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) + ∷ invPa1 + + let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms + (finalTrm0 ++ symPathTerms finalTrm1) + + pure (finalTrm , dbgInfo) + + +groupoidSolverMain : Bool → R.Term → R.Term → R.Term → R.TC Unit +groupoidSolverMain debugFlag inp inpF hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag inp inpF hole' + R.catchTC + (R.unify hole solution) + (R.typeError msg) + +groupoidSolverMain' : Bool → R.Term → R.TC Unit +groupoidSolverMain' debugFlag hole = do + + let inpF = R.con (quote [fn]) [] + let inp = R.con (quote [Tω]) [] + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + hole' ← R.withNormalisation true $ R.checkType hole ty + R.commitTC + (solution , msg) ← groupoidSolverTerm debugFlag inp inpF hole' + -- R.typeError [ solution ]ₑ + R.catchTC + (R.unify hole solution) + (R.typeError msg) + + +-- -- squareSolverMain : Bool → R.Term → R.Term → R.TC Unit +-- -- squareSolverMain debugFlag inp hole = do +-- -- ty ← R.inferType hole >>= wait-for-type +-- -- hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta + +-- -- (solution , msg) ← groupoidSolverTerm debugFlag inp hole' +-- -- R.catchTC +-- -- (R.unify hole' solution) +-- -- (R.typeError msg) + +-- -- R.catchTC +-- -- (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) +-- -- (R.typeError [ R.strErr "xxx" ] ) +-- -- where +-- -- extractMeta : R.Term → R.TC R.Term +-- -- extractMeta (R.def _ tl) = matchVarg tl +-- -- extractMeta t = +-- -- R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) + +macro + solveGroupoidDebug : R.Term → R.Term → R.Term → R.TC Unit + solveGroupoidDebug = groupoidSolverMain true + + ≡! : R.Term → R.TC Unit + ≡! = groupoidSolverMain' true + +-- -- solveSquareDebug : R.Term → R.TC Unit +-- -- solveSquareDebug = squareSolverMain false + +-- -- solveGroupoid : R.Term → R.TC Unit +-- -- solveGroupoid = groupoidSolverMain true + +-- -- solveSquare : R.Term → R.TC Unit +-- -- solveSquare = squareSolverMain false + + +module SimpleTest where + + module _ (n : ℕ) where + + data AA : Type where + x y z w : AA + p p' : x ≡ y + q : y ≡ z + q' : z ≡ y + r : z ≡ w + + + + pA pB : Path (AA 3) (x) (w) + pA = ((refl ∙ p) ∙ q) ∙ r + pB = ((refl ∙ p) ∙ q) ∙ r + + + pA=pB : Path (Path (AA 3) x y) + (refl ∙ p) + (p) + pA=pB = ≡! + + +module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where + + pA pB pC : x ≡ w + pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + + pA=pB : pA ≡ pB + pA=pB = ≡! + + +module Examples' (A : ℕ → Type ℓ) + (x y z w : ∀ {n} → A n) + (p p' : ∀ {n} → x {n} ≡ y) + (q : ∀ {n} → y {n} ≡ z) + (q' : ∀ {n} → z {n} ≡ y) + (r : ∀ {n} → w {n} ≡ z) where + + pA pB : Path (A 3) x w + pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + + pA=pB : pA ≡ pB + pA=pB = ≡! + + +module Examplesℕ' where + + data AA : (n : ℕ) → Type where + x y z w : ∀ {n} → AA n + p p' : ∀ {n} → x {n} ≡ y + q : ∀ {n} → y {n} ≡ z + q' : ∀ {n} → z {n} ≡ y + r : ∀ {n} → w {n} ≡ z + + + + pA pB : Path (AA 3) (x) (w) + pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) + pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) + ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl + + pA=pB : pA ≡ pB + pA=pB = ≡! + + +module Examplesℕ'' where + + module _ (n : ℕ) where + + data AA : Type where + x y z w : AA + p p' : x ≡ y + q : y ≡ z + q' : z ≡ y + r : w ≡ z + + + + pA pB : Path (AA 3) (x) (w) + pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) + pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) + ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl + + + pA=pB : pA ≡ pB + pA=pB = ≡! + + +-- module Examplesℕ where + +-- data AA (n : ℕ) : Type where +-- x y z w : AA n +-- p p' : x ≡ y +-- q : y ≡ z +-- q' : z ≡ y +-- r : w ≡ z + + + +-- pA pB : Path (AA 3) (x) (w) +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) +-- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) +-- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl +-- -- -- -- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + +-- pA=pB : pA ≡ pB +-- pA=pB = ≡! + + + + +-- -- -- -- -- -- -- -- module Examples2 (A B : Type ℓ) (x y z : B) (w : A) (f g : B → A) +-- -- -- -- -- -- -- -- (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ f z) where + +-- -- -- -- -- -- -- -- pA pB : f x ≡ w +-- -- -- -- -- -- -- -- pA = cong f (p ∙∙ refl ∙∙ q) ∙ sym r +-- -- -- -- -- -- -- -- pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl + + +-- -- -- -- -- -- -- -- pA=pB : pA ≡ pB +-- -- -- -- -- -- -- -- pA=pB = solveGroupoidDebug (B ∷Tω A ∷Tω [Tω]) (g ∷fn f ∷fn [fn]) + diff --git a/Cubical/Tactics/PathSolver/Solver5.agda b/Cubical/Tactics/PathSolver/Solver5.agda new file mode 100644 index 0000000000..db236cde76 --- /dev/null +++ b/Cubical/Tactics/PathSolver/Solver5.agda @@ -0,0 +1,876 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Solver5 where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool as 𝟚 hiding (_≤_) +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.List as Li +open import Cubical.Data.Maybe as Mb + + +open import Cubical.HITs.Interval + +open import Cubical.Relation.Nullary + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.WildCatSolver.Reflection +open import Cubical.Tactics.Reflection +open import Agda.Builtin.String + + +private + variable + ℓ ℓ' : Level + A A' : Type ℓ + + +infixr 5 _∷Tω_ + +infixr 5 _∷fn_ + + +data [Typeₙ] : Typeω where + [Tω] : [Typeₙ] + _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] + +data [Fns] : Typeω where + [fn] : [Fns] + _∷fn_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] + + +reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) +reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] _ = + R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ + + +-- macro +-- test-refl[T]macro : R.Term → R.Term → R.TC Unit +-- test-refl[T]macro inp hole = do +-- l ← reflected[Typeₙ]→[reflectedTy] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[T] : Unit +-- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) + + +reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) +reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ + +-- macro +-- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit +-- test-refl[Fn]macro inp hole = do +-- l ← reflected[Fns]→[reflectedFn] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[Fn] : Unit +-- test-refl[Fn] = test-refl[Fn]macro +-- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) + + +maxℓ : [Typeₙ] → Level +maxℓ [Tω] = ℓ-zero +maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) + +lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) +lookupTₙ [Tω] x = ⊥* +lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ +lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) + +ℓAt : (Ts : [Typeₙ]) → ℕ → Level +ℓAt [Tω] x = ℓ-zero +ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ +ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x + +TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) +TyAt [Tω] k = ⊥* +TyAt (x ∷Tω Ts) zero = x +TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k + +cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k +cast↓ [Tω] k () +cast↓ (x₁ ∷Tω Ts) zero x = lower x +cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) + +cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y +cast↓Inj {[Tω]} {A = A} {()} +cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift +cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = + cong lift (cast↓Inj {[T] = [T]} {A = A} p) + +cast↓Inj-sec : ∀ {[T] A x y} p → + cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p +cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl +cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = + cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) + +cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → + cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) + ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) + + +cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ +cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = + cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ + + +cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k +cast↑ [Tω] k () +cast↑ (x₁ ∷Tω Ts) zero x = lift x +cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x + +-- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) +-- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t + + +-- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) +-- Ts-arr' [] x x₁ = Unit +-- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) +-- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} +-- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} +-- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = +-- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) + +-- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → +-- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t +-- Ts-arr' Ts s t x x₁ = {!!} + + + +data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) + compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) + congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') + → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) + + +module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} + (p : Path {ℓ} A x y) where + + -- pathCaseCongTest : PathCase λ i → f (p i) + -- pathCaseCongTest = congCase f {!!} + + +infixl 5 _fp∷_ +infixl 5 _fp++_ + +data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where + [fp] : ∀ {x b} → FlatPath A b x x + _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z + _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y + +magicNumber = 100 + +mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) +mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing +mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w + where + w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) + w (R.con (quote [fp]) _ , _) = R.returnTC nothing + w (xs'@(R.con (quote _fp∷_) tl') , y) = + match2Vargs tl' >>= λ (xs , x) → + R.catchTC + (R.noConstraints $ R.unify + (R.def (quote sym) (x v∷ [])) y + >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) + >>= λ (xs* , xs*↑) → + R.returnTC + (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) + , xs*↑ ) + )) + (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) + ((xs* v∷ y v∷ [])))) + <$> mb-invol n xs') + w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) +mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) + +mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol' = mb-invol magicNumber + + +redList : ℕ → R.Term → R.TC (List R.Term) +redList = h + where + h : ℕ → R.Term → R.TC (List R.Term) + h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) + h (suc k) t = + (mb-invol' t) >>= + Mb.rec + (R.returnTC []) + λ (t' , t'↓) → do + p' ← h k t'↓ + R.returnTC (t' ∷ p') + + +redList' : R.Term → R.TC (List R.Term) +redList' = redList magicNumber + + +pattern fp[_] x = [fp] fp∷ x + +FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ +FP⟦ [fp] ⟧ = refl +FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ + +-- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ +-- FP⟦ [] ⟧t = refl +-- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ +-- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ + + +fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ +fpF→T [fp] = [fp] +fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ + +fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ +fpT→F _ [fp] = [fp] +fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ +fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ +fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ + +fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → + FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ +fpT≡F [fp] = refl +fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x +fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = + hcomp + (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ k ∧ ~ i) + }) + (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ i ∧ k) + }) + (fpT≡F fp i j)) + +_fp++_ : ∀ {x y z} + → FlatPath A false x y + → FlatPath A false y z + → FlatPath A false x z +x fp++ [fp] = x +x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ + +fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) + (fp' : FlatPath A false a₁ a₂) + → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ +fp++∙ fp [fp] = sym (rUnit _) +fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') + +module PE ([T] : [Typeₙ]) where + + data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where + reflExpr : ∀ {A x} → PathExpr A x x + atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y + compExpr : ∀ {A x y z w} + → PathExpr A x y → PathExpr A y z → PathExpr A z w + → PathExpr A x w + congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') + → PathExpr A x y + → PathExpr A' (f x) (f y) + + + PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → + (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) + PE⟦ reflExpr ⟧ = refl + PE⟦ atomExpr p ⟧ = cong _ p + PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ + PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) + + cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) + → PathExpr A₁ a₀ a₁ + + → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) + (cast↓ [T] A (f a₁)) + cong-flat f reflExpr = [fp] + cong-flat f (atomExpr p) = fp[ cong _ p ] + cong-flat f (compExpr x x₁ x₂) = + cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ + cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x + + flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ + → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) + flat⟦ reflExpr ⟧ = [fp] + flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] + flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ + flat⟦ congExpr f x ⟧ = cong-flat f x + + + cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → + (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → + (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ + FP⟦ cong-flat f pe ⟧ + cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) + cong-flat≡ (atomExpr p) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ + cong-flat≡ (compExpr pe pe₁ pe₂) f = + (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ + cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ + (λ i → doubleCompPath-elim + (cong-flat≡ pe f i) + (cong-flat≡ pe₁ f i) + (cong-flat≡ pe₂ f i) i)) + ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) + (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) + ∙∙ fp++∙ _ (cong-flat f pe₂) + cong-flat≡ (congExpr f₁ pe) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) + + pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → + PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ + pe≡flat reflExpr = refl + pe≡flat (atomExpr p) = lUnit _ + pe≡flat (compExpr pe pe₁ pe₂) = + (λ i → doubleCompPath-elim + (pe≡flat pe i) + (pe≡flat pe₁ i) + (pe≡flat pe₂ i) i) + ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) + ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ + + pe≡flat (congExpr f pe) = cong-flat≡ pe f + + +module PathTrm (A B : Type ℓ) where + data PathTrm : Type ℓ where + reflTrm : PathTrm + atomTrm : A → PathTrm + compTrm : PathTrm → PathTrm → PathTrm → PathTrm + congTrm : B → PathTrm → PathTrm + + module showPathTrm (showA : A → String) (showB : B → String) where + showPT : PathTrm → String + showPT reflTrm = "refl" + showPT (atomTrm x) = showA x + showPT (compTrm x x₁ x₂) = + "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" + showPT (congTrm x x₁) = + "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" + + +module _ {ℓ ℓ'} + {A B : Type ℓ} + {A' B' : Type ℓ'} + (f : A → R.TC A') + (g : B → R.TC B') where + open PathTrm + mmapPT : PathTrm A B → R.TC $ PathTrm A' B' + mmapPT reflTrm = pure reflTrm + mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ + mmapPT (compTrm x x₁ x₂) = + ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ + mmapPT (congTrm x x₁) = + ⦇ congTrm (g x) (mmapPT x₁) ⦈ + + +module RTrm = PathTrm R.Term R.Term +module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) +module StrTrm = PathTrm String String + +showRTrmTC : RTrm.PathTrm → R.TC String +showRTrmTC = + mmapPT + (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +showRTrmTC' : RTrm'.PathTrm → R.TC String +showRTrmTC' = + let q = λ (k , t) → + R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) + in mmapPT + q q + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +module _ ([T] : [Typeₙ]) where + reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) + reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} + + + atomExpr' : ∀ A {x y} → (p : x ≡ y) → + PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) + + compExpr' : ∀ A {x y z w} + → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w + → PE.PathExpr [T] A x w + compExpr' A = PE.compExpr {[T] = [T]} {A = A} + + congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') + → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) + (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) + congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} + (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x + +getEnd : ∀ {x y : A} → x ≡ y → A +getEnd p = p i0 + +module tryPathE --([T] : [Typeₙ]) + (TC[T]trm : R.Term) + ([TC[T]trm] [fns] : List R.Term) where + + inferA : R.Term → R.TC ℕ + inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + + inferA' : R.Term → R.TC R.Term + inferA' = inferA >=> R.quoteTC + + + + try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + + + tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + let x₀ = R.def (quote getEnd) v[ t ] + A ← inferA' t + ⦇ (RTrm.reflTrm , R.def (quote reflExpr') + (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ + + tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (t1 , t1') ← (try≡ k t1) + (t2 , t2') ← (try≡ k t2) + (t3 , t3') ← (try≡ k t3) + A ← inferA' t + pure (RTrm.compTrm t1 t2 t3 , + R.def (quote compExpr') + (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t*) ← h tm + A ← inferA' t* + A' ← inferA' t + (t0 , t0') ← (try≡ k t*) + pure (RTrm.congTrm f t0 , + R.def (quote congExpr') + (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) + + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + + atom : R.Term → R.TC (RTrm.PathTrm × R.Term) + atom x = do + A ← inferA' x + pure (RTrm.atomTrm x , R.def (quote atomExpr') + (TC[T]trm v∷ A v∷ x v∷ [])) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) + + + +module tryTermE --([T] : [Typeₙ]) + --(TC[T]trm : R.Term) + ([TC[T]trm] [fns] : List R.Term) where + + try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) + + + tryRefl : R.Term → R.TC (RTrm'.PathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + ⦇ RTrm'.reflTrm ⦈ + + tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ + + where + + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t) ← h tm + ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) + + where + + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] + + + + atom : R.Term → R.TC (RTrm'.PathTrm) + atom x = do + k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + R.returnTC (RTrm'.atomTrm (k , x)) + + + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) + + + +-- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) +-- funTypeView = {!!} + +groupoidSolverTerm : Bool → R.Term → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag inp₀ inpF ty = do + + -- inp₀ ← wait-for-type inp₀' + + -- R.typeError (niceAtomList [) + (A' , (t0' , t1')) ← wait-for-type ty >>= (get-boundaryWithDom ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) + + -- A ← wait-for-type A' + -- t0 ← wait-for-type t0' + -- t1 ← wait-for-type t1' + + let t0 = t0' + let t1 = t1' + + (AA , _) ← get-boundaryWithDom A' >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary2" ]) + (λ x → R.returnTC x) + + let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) + -- (R.typeError [ inp ]ₑ) + + + + [t] ← reflected[Typeₙ]→[reflectedTy] inp + [f] ← reflected[Fns]→[reflectedFn] inpF + + (r0 , r0') ← tryPathE.try≡ inp [t] [f] 100 t0 + (r1 , r1') ← tryPathE.try≡ inp [t] [f] 100 t1 + + + + -- (aL' , (_ , e0)) ← unMapAtoms [] r0' + -- (aL , (_ , e1)) ← unMapAtoms aL' r1' + -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + show0 ← showRTrmTC r0 + show1 ← showRTrmTC r1 + + red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' + red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' + + + let invPa0 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red0 + invPa1 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red1 + + let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl + ("LHS: ") ∷ₑ show0 ∷nl + ("RHS: ") ∷ₑ show1 ∷nl + (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) + -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) + + -- ∷ (R.strErr "\n") + -- (niceAtomList aL) + + let finalTrm0 = + just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) + ∷ invPa0 + + finalTrm1 = + just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) + ∷ invPa1 + + let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms + (finalTrm0 ++ symPathTerms finalTrm1) + + pure (finalTrm , dbgInfo) + + +groupoidSolverMain : Bool → R.Term → R.Term → R.Term → R.TC Unit +groupoidSolverMain debugFlag inp inpF hole = do + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + -- hole' ← R.withNormalisation true $ R.checkType hole ty + (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty + R.catchTC + (R.unify hole solution) + (R.typeError msg) + +groupoidSolverMain' : Bool → R.Term → R.TC Unit +groupoidSolverMain' debugFlag hole = do + + let inpF = R.con (quote [fn]) [] + let inp = R.con (quote [Tω]) [] + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type + -- hole' ← R.withNormalisation true $ R.checkType hole ty + + (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty + -- R.typeError [ solution ]ₑ + R.catchTC + (R.unify hole solution) + (R.typeError msg) + + +groupoidSolverMainTerm' : Bool → R.Term → R.Term → R.TC Unit +groupoidSolverMainTerm' debugFlag ty hole = do + + let inpF = R.con (quote [fn]) [] + let inp = R.con (quote [Tω]) [] + + -- hole' ← R.withNormalisation true $ R.checkType hole ty + -- R.commitTC + (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty + solutionR ← R.quoteTC solution + R.catchTC + (R.unify hole solutionR) + (R.typeError msg) + + + +-- -- squareSolverMain : Bool → R.Term → R.Term → R.TC Unit +-- -- squareSolverMain debugFlag inp hole = do +-- -- ty ← R.inferType hole >>= wait-for-type +-- -- hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta + +-- -- (solution , msg) ← groupoidSolverTerm debugFlag inp hole' +-- -- R.catchTC +-- -- (R.unify hole' solution) +-- -- (R.typeError msg) + +-- -- R.catchTC +-- -- (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) +-- -- (R.typeError [ R.strErr "xxx" ] ) +-- -- where +-- -- extractMeta : R.Term → R.TC R.Term +-- -- extractMeta (R.def _ tl) = matchVarg tl +-- -- extractMeta t = +-- -- R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) + +macro + solveGroupoidDebug : R.Term → R.Term → R.Term → R.TC Unit + solveGroupoidDebug = groupoidSolverMain true + + ≡! : R.Term → R.TC Unit + ≡! = groupoidSolverMain' true + + ≡!trm : R.Term → R.Term → R.TC Unit + ≡!trm = groupoidSolverMainTerm' true + + +-- -- solveSquareDebug : R.Term → R.TC Unit +-- -- solveSquareDebug = squareSolverMain false + +-- -- solveGroupoid : R.Term → R.TC Unit +-- -- solveGroupoid = groupoidSolverMain true + +-- -- solveSquare : R.Term → R.TC Unit +-- -- solveSquare = squareSolverMain false + + +-- module SimpleTest where + +-- module _ (n : ℕ) where + +-- data AA : Type where +-- x y z w : AA +-- p p' : x ≡ y +-- q : y ≡ z +-- q' : z ≡ y +-- r : z ≡ w + + + +-- pA pB : Path (AA 3) (x) (w) +-- pA = ((refl ∙ p) ∙ q) ∙ r +-- pB = ((refl ∙ p) ∙ q) ∙ r + + +-- pA=pB : Path (Path (AA 3) x y) +-- (refl ∙ p) +-- (p) +-- pA=pB = ≡! + + +module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where + + pA pB pC : x ≡ w + pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + + pA=pB : pA ≡ pB + pA=pB = ≡! + + +-- module Examples' (A : ℕ → Type ℓ) +-- (x y z w : ∀ {n} → A n) +-- (p p' : ∀ {n} → x {n} ≡ y) +-- (q : ∀ {n} → y {n} ≡ z) +-- (q' : ∀ {n} → z {n} ≡ y) +-- (r : ∀ {n} → w {n} ≡ z) where + +-- pA pB : Path (A 3) x w +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym r +-- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl + +-- pA=pB : pA ≡ pB +-- pA=pB = ≡! + + +-- module Examplesℕ' where + +-- data AA : (n : ℕ) → Type where +-- x y z w : ∀ {n} → AA n +-- p p' : ∀ {n} → x {n} ≡ y +-- q : ∀ {n} → y {n} ≡ z +-- q' : ∀ {n} → z {n} ≡ y +-- r : ∀ {n} → w {n} ≡ z + + + +-- pA pB : Path (AA 3) (x) (w) +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) +-- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) +-- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl + +-- pA=pB : pA ≡ pB +-- pA=pB = ≡! + + + + +-- module Examplesℕ where + +-- data AA (n : ℕ) : Type where +-- x y z w : AA n +-- p p' : x ≡ y +-- q : y ≡ z +-- q' : z ≡ y +-- r : w ≡ z + + + +-- pA pB : Path (AA 3) (x) (w) +-- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) +-- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) +-- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl +-- -- -- -- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r + +-- pA=pB : R.Term +-- pA=pB = {!≡!trm (pA ≡ pB)!} + + + + +-- -- -- -- -- -- -- -- -- -- module Examples2 (A B : Type ℓ) (x y z : B) (w : A) (f g : B → A) +-- -- -- -- -- -- -- -- -- -- (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ f z) where + +-- -- -- -- -- -- -- -- -- -- pA pB : f x ≡ w +-- -- -- -- -- -- -- -- -- -- pA = cong f (p ∙∙ refl ∙∙ q) ∙ sym r +-- -- -- -- -- -- -- -- -- -- pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl + + +-- -- -- -- -- -- -- -- -- -- pA=pB : pA ≡ pB +-- -- -- -- -- -- -- -- -- -- pA=pB = solveGroupoidDebug (B ∷Tω A ∷Tω [Tω]) (g ∷fn f ∷fn [fn]) + diff --git a/Cubical/Tactics/Reflection.agda b/Cubical/Tactics/Reflection.agda index 802ffd21e9..f31e4d5134 100644 --- a/Cubical/Tactics/Reflection.agda +++ b/Cubical/Tactics/Reflection.agda @@ -90,6 +90,12 @@ get-boundary tm = unapply-path tm >>= λ where (just (_ , x , y)) → returnTC (just (x , y)) nothing → returnTC nothing +get-boundaryWithDom : Term → TC (Maybe (Term × (Term × Term))) +get-boundaryWithDom tm = unapply-path tm >>= λ where + (just (A , x , y)) → returnTC (just (A , (x , y))) + nothing → returnTC nothing + + equation-solver : List Name → (Term -> Term -> TC Term) → Bool → Term → TC Unit equation-solver don't-Reduce mk-call debug hole = withNormalisation false ( diff --git a/Cubical/Tactics/WildCatSolver/Reflection.agda b/Cubical/Tactics/WildCatSolver/Reflection.agda index 288e3c99de..10cf2ca65f 100644 --- a/Cubical/Tactics/WildCatSolver/Reflection.agda +++ b/Cubical/Tactics/WildCatSolver/Reflection.agda @@ -221,6 +221,13 @@ foldlM f b (x ∷ xs) = f b x >>= (flip (foldlM f)) xs R∙ : R.Term → R.Term → R.Term R∙ x y = R.def (quote _∙_) (x v∷ y v∷ [] ) +R∙∙ : R.Term → R.Term → R.Term → R.Term +R∙∙ x y z = R.def (quote _∙∙_∙∙_) (x v∷ y v∷ z v∷ [] ) + +Rsym : R.Term → R.Term +Rsym x = R.def (quote sym) (x v∷ [] ) + + Rrefl : R.Term Rrefl = R.def (quote refl) [] From c1376ac3887053c90162b647c97c4cd704b3a6be Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Tue, 2 Apr 2024 03:53:33 +0200 Subject: [PATCH 11/20] cleanup --- Cubical/Foundations/GroupoidLaws.agda | 2 +- Cubical/Tactics/PathSolver/Example.agda | 76 +- .../Tactics/PathSolver/ReflectionTest.agda | 50 - Cubical/Tactics/PathSolver/Solver.agda | 1036 +++++++++++------ Cubical/Tactics/PathSolver/Solver2.agda | 647 ---------- Cubical/Tactics/PathSolver/Solver3.agda | 806 ------------- Cubical/Tactics/PathSolver/Solver4.agda | 877 -------------- Cubical/Tactics/PathSolver/Solver5.agda | 876 -------------- Cubical/Tactics/Reflection.agda | 54 +- 9 files changed, 762 insertions(+), 3662 deletions(-) delete mode 100644 Cubical/Tactics/PathSolver/ReflectionTest.agda delete mode 100644 Cubical/Tactics/PathSolver/Solver2.agda delete mode 100644 Cubical/Tactics/PathSolver/Solver3.agda delete mode 100644 Cubical/Tactics/PathSolver/Solver4.agda delete mode 100644 Cubical/Tactics/PathSolver/Solver5.agda diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index c632117408..9b36dc6fdc 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -84,7 +84,7 @@ midCancel p q r j = (λ i → p (i ∧ j)) ∙∙ doubleCompPath-filler p refl (sym q) (~ j) ∙∙ λ i → hcomp (λ k → λ {(i = i0) → q (~ k ∨ j) ; (i = i1) → r (~ k) ; (j = i1) → r (~ i ∨ ~ k) }) (r i1) - + lCancel : (p : x ≡ y) → p ⁻¹ ∙ p ≡ refl lCancel p = rCancel (p ⁻¹) diff --git a/Cubical/Tactics/PathSolver/Example.agda b/Cubical/Tactics/PathSolver/Example.agda index 99c852fba9..3cdb601bd2 100644 --- a/Cubical/Tactics/PathSolver/Example.agda +++ b/Cubical/Tactics/PathSolver/Example.agda @@ -6,58 +6,82 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path open import Cubical.Foundations.Structure +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + open import Cubical.Tactics.PathSolver.Solver private variable - ℓ ℓ' : Level + ℓ : Level + A B C : Type ℓ -module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where +module Examples (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where pA pB pC : x ≡ w - pA = (p ∙∙ refl ∙∙ q) ∙ sym r + pA = (p ∙∙ refl ∙∙ q) ∙' sym r pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r pA=pB : pA ≡ pB - pA=pB = solveGroupoidDebug + pA=pB = ≡! pB=pC : pB ≡ pC - pB=pC = solveGroupoidDebug + pB=pC = ≡! pA=pC : pA ≡ pC - pA=pC = solveGroupoidDebug + pA=pC = ≡! - pA=pB∙pB=pC-≡-pA=pC : pA=pB ∙ pB=pC ≡ pA=pC - pA=pB∙pB=pC-≡-pA=pC = midCancel _ _ _ + sqTest : Square p (sym r ∙ refl) (p ∙ q) (q ∙ sym r) + sqTest = sq! - sqTest : Square p (sym r ∙ refl) (p ∙ q) (q ∙ sym r) - sqTest = solveSquareDebug +module Examples2 (x y z : B) (w : A) (f g : B → A) + (p p' : x ≡ y) (q : y ≡ z) (r : w ≡ f z) where -module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where + pA pB : f x ≡ w + pA = cong f (p' ∙ sym p') ∙' cong f (p ∙∙ refl ∙∙ q) ∙ sym r + pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl + + + pA=pB : pA ≡ pB + pA=pB = ≡!cong f - open import Cubical.Foundations.Equiv - SqCompEquiv : (Square p p q q) ≃ (p ∙ q ≡ q ∙ p) - SqCompEquiv = 2-cylinder-from-square.Sq≃Sq' refl solveSquareDebug +module Examples3 (x y z : B) (w : A) (f : B → A) + (p p' : x ≡ y) (q q' : y ≡ z) (r r' : f z ≡ w) where + squares≃Example : (Square (cong f p') (cong f q ∙ r) (cong f p) (cong f q' ∙ r')) + ≃ (cong f (sym (p ∙ q) ∙∙ p' ∙∙ q') ≡ r ∙ sym r') -open import Cubical.Algebra.Group + squares≃Example = + 2-cylinder-from-square.Sq≃Sq' + (cong f ((p ∙ q))) + (≡!cong f) -module EM₁-Example (G : Group ℓ) (a b c : fst G) where - open GroupStr (snd G) - data EM₁ : Type ℓ where - embase : EM₁ - emloop : ⟨ G ⟩ → embase ≡ embase - emcomp : (g h : ⟨ G ⟩ ) → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) +module Examples4 (f : B → A) (g : C → B) + (x y : B) (z w v : C) + (p : x ≡ y) (q : y ≡ g z) (r : z ≡ w) (u : w ≡ v) where - double-emcomp : Square {A = EM₁} - (emloop b) (emloop ((a · b) · c)) - (sym (emloop a)) (emloop c) - double-emcomp = fst (2-cylinder-from-square.Sq≃Sq' - (emloop a) solveSquareDebug) (emcomp a b ∙v emcomp (a · b) c) + sqE : (cong f p ∙∙ cong f q ∙∙ cong (f ∘ g) r) ∙ (cong (f ∘ g) u) + ≡ (cong f (p ∙ q)) ∙ (cong f (cong g r ∙ cong g u )) + sqE = ≡!cong (f ∷ g ∷ [fn]) + + + sqE' : Square (cong f p ∙∙ cong f q ∙∙ cong (f ∘ g) r) + (cong f (cong g r ∙ cong g u )) + (cong f (p ∙ q)) + (cong (f ∘ g) u) + sqE' = sq!cong (f ∷ g ∷ [fn]) + + +module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where + + open import Cubical.Foundations.Equiv + + SqCompEquiv : (Square p p q q) ≃ (p ∙ q ≡ q ∙ p) + SqCompEquiv = 2-cylinder-from-square.Sq≃Sq' refl ≡! diff --git a/Cubical/Tactics/PathSolver/ReflectionTest.agda b/Cubical/Tactics/PathSolver/ReflectionTest.agda deleted file mode 100644 index b32b5c44b0..0000000000 --- a/Cubical/Tactics/PathSolver/ReflectionTest.agda +++ /dev/null @@ -1,50 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.ReflectionTest where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - - -module Test1 where - - module _ (n : ℕ) where - data A : Type where - a aa : A - p : a ≡ aa - - macro - unquoteA : R.Term → R.TC Unit - unquoteA hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - final ← R.checkType (R.def (quote refl) []) ty >>= R.normalise - R.unify hole final - -- (R∙ (R.def (quote refl) []) (R.def (quote refl) [])) - - a' : Path (Path (A 3) a aa) (refl ∙ p) (refl ∙ p) - a' = unquoteA diff --git a/Cubical/Tactics/PathSolver/Solver.agda b/Cubical/Tactics/PathSolver/Solver.agda index 324669b52c..b040b784fe 100644 --- a/Cubical/Tactics/PathSolver/Solver.agda +++ b/Cubical/Tactics/PathSolver/Solver.agda @@ -13,6 +13,7 @@ open import Cubical.Foundations.Path open import Cubical.Data.Bool as 𝟚 hiding (_≤_) open import Cubical.Data.Nat as ℕ hiding (_·_) open import Cubical.Data.Unit +open import Cubical.Data.Empty open import Cubical.Data.Sigma open import Cubical.Data.List as Li open import Cubical.Data.Maybe as Mb @@ -32,451 +33,740 @@ open import Agda.Builtin.String private variable ℓ ℓ' : Level + A A' : Type ℓ + + +infixr 5 _∷Tω_ + +infixr 5 _∷_ + + +typeView : R.Term → R.TC (Maybe R.Term) +typeView t = + ⦇ just (R.checkType t (R.def (quote Type) v[ R.unknown ])) ⦈ <|> ⦇ nothing ⦈ + + + +-- macro +-- tvTest : R.Term → R.Term → R.TC Unit +-- tvTest ty? hole = +-- typeView ty? >>= +-- Mb.caseMaybe +-- (R.typeError [ "not a type!" ]ₑ) +-- (R.typeError [ "is type!" ]ₑ) + + +data [Typeₙ] : Typeω where + [Tω] : [Typeₙ] + _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] + +data [Fns] : Typeω where + [fn] : [Fns] + _∷_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] + + +gatherCodoms : R.Term → R.TC (List R.Term) +gatherCodoms t@(R.con (quote [fn]) _) = pure [] +gatherCodoms t@(R.con (quote [Fns]._∷_) (_ h∷ _ h∷ a' h∷ _ h∷ _ v∷ xs v∷ [])) = + ⦇ ⦇ a' ⦈ ∷ gatherCodoms xs ⦈ +gatherCodoms t = R.typeError $ "gatherCodoms - FAIL : " ∷ₑ [ t ]ₑ + +inferTypesFromFns : R.Term → R.TC (List R.Term) +inferTypesFromFns = + (wait-for-term >=> gatherCodoms) >=> uniqeAtoms + +reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) +reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ +reflected[Typeₙ]→[reflectedTy] _ = R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ + -- typeView x >>= + -- Mb.rec + -- (R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ) + -- (pure ∘ [_]) + +[reflectedTy]→reflected[Typeₙ] : List R.Term → R.Term +[reflectedTy]→reflected[Typeₙ] [] = (R.con (quote [Tω]) []) +[reflectedTy]→reflected[Typeₙ] (x ∷ xs) = + (R.con (quote _∷Tω_) (x v∷ [reflectedTy]→reflected[Typeₙ] xs v∷ [])) + +castTo[Typeₙ] : R.Term → R.TC (R.Term) +castTo[Typeₙ] t@(R.con (quote [Tω]) _) = pure t +castTo[Typeₙ] t@(R.con (quote _∷Tω_) _) = pure t +castTo[Typeₙ] x = + typeView x >>= + Mb.rec + (R.typeError [ "castTo[Typeₙ] - FAIL" ]ₑ) + (λ t → pure (R.con (quote _∷Tω_) (t v∷ (R.con (quote [Tω]) []) v∷ []))) + +castTo[Fns] : R.Term → R.TC (R.Term) +castTo[Fns] t@(R.con (quote [fn]) _) = pure t +castTo[Fns] t@(R.con (quote [Fns]._∷_) _) = pure t +castTo[Fns] x = + R.checkType + (R.con (quote [Fns]._∷_) (x v∷ (R.con (quote [fn]) []) v∷ [])) + (R.def (quote [Fns]) []) + -- funView x >>= + -- Mb.rec + -- (R.typeError [ "castTo[Fns] - FAIL" ]ₑ) + -- (λ t → pure (R.con (quote [Fns]._∷_) (t v∷ (R.con (quote [fn]) []) v∷ []))) + + + +-- macro +-- test-refl[T]macro : R.Term → R.Term → R.TC Unit +-- test-refl[T]macro inp hole = do +-- l ← reflected[Typeₙ]→[reflectedTy] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[T] : Unit +-- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) + + +reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) +reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] +reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (x v∷ x₁ v∷ [])) = + (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ +reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ + +-- macro +-- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit +-- test-refl[Fn]macro inp hole = do +-- l ← reflected[Fns]→[reflectedFn] inp +-- R.typeError (niceAtomList l) + +-- module _ (ℓ ℓ' : Level) where +-- test-refl[Fn] : Unit +-- test-refl[Fn] = test-refl[Fn]macro +-- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) + + +maxℓ : [Typeₙ] → Level +maxℓ [Tω] = ℓ-zero +maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) + +lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) +lookupTₙ [Tω] x = ⊥* +lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ +lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) + +ℓAt : (Ts : [Typeₙ]) → ℕ → Level +ℓAt [Tω] x = ℓ-zero +ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ +ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x + +TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) +TyAt [Tω] k = ⊥* +TyAt (x ∷Tω Ts) zero = x +TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k + +cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k +cast↓ [Tω] k () +cast↓ (x₁ ∷Tω Ts) zero x = lower x +cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) + +cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y +cast↓Inj {[Tω]} {A = A} {()} +cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift +cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = + cong lift (cast↓Inj {[T] = [T]} {A = A} p) + +cast↓Inj-sec : ∀ {[T] A x y} p → + cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p +cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl +cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = + cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) + +cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → + cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) + ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) + + +cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ +cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = + cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ + + +cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k +cast↑ [Tω] k () +cast↑ (x₁ ∷Tω Ts) zero x = lift x +cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x + +-- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) +-- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t +-- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) +-- Ts-arr' [] x x₁ = Unit +-- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) +-- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} +-- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} +-- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = +-- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) -module _ {A : Type ℓ} where - data PathCase : {a₀ a₁ : A} → a₀ ≡ a₁ → Type ℓ where - reflCase : ∀ {x} → PathCase (refl {x = x}) - compCase : ∀ {x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase (p ∙∙ q ∙∙ r) +-- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → +-- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t +-- Ts-arr' Ts s t x x₁ = {!!} -module PT {A : Type ℓ} (_∼_ : A → A → Type ℓ') (∼refl : ∀ {x} → x ∼ x) - (_∼∙_ : ∀ {x y z} → x ∼ y → y ∼ z → x ∼ z) - (_∼∙∙_∼∙∙_ : ∀ {x y z w} → x ∼ y → y ∼ z → z ∼ w → x ∼ w) - (∼doubleCompPath-elim : ∀ {x y z w} → - (p : x ∼ y) → (q : y ∼ z) → (r : z ∼ w) → (p ∼∙∙ q ∼∙∙ r) ≡ (p ∼∙ q) ∼∙ r) - (∼assoc : ∀ {x y z w} → (p : x ∼ y) (q : y ∼ z) (r : z ∼ w) → p ∼∙ (q ∼∙ r) ≡ (p ∼∙ q) ∼∙ r) - where +data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where + reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) + compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) + congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') + → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - data IsPathTrm : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - isReflTrm : ∀ {x} → IsPathTrm (∼refl {x = x}) - isAtomTrm : ∀ {x y} → (p : x ∼ y) → IsPathTrm p - isCompTrm : ∀ {x y z w : _} → {p : x ∼ y} {q : y ∼ z} {r : z ∼ w} - → IsPathTrm p → IsPathTrm q → IsPathTrm r → IsPathTrm (p ∼∙∙ q ∼∙∙ r) +module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} + (p : Path {ℓ} A x y) where + -- pathCaseCongTest : PathCase λ i → f (p i) + -- pathCaseCongTest = congCase f {!!} +infixl 5 _fp∷_ +infixl 5 _fp++_ - data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) - isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p - isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} - → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) +data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where + [fp] : ∀ {x b} → FlatPath A b x x + _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z + _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y - data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) - consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) - - data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where - nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x - consTrmInvol : ∀ {x y z : _} → - IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z - involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → - IsPathTrmInvol x y - - - module show (showA : {a₀ a₁ : A} → (p : a₀ ∼ a₁) → String) where - showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String - showIPT isReflTrm = "refl" - showIPT (isAtomTrm x) = showA x - showIPT (isCompTrm isReflTrm x₁ x₂) = - "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" - showIPT (isCompTrm x x₁ isReflTrm) = - "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" - showIPT (isCompTrm x x₁ x₂) = - "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" - - showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String - - showIPTD nilTrmDeasRefl = "refl" - showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q - - showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String - showIPTI nilTrmInvolRefl = "refl" - showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q - showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" +magicNumber = 100 +mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) +mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing +mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w + where + w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) + w (R.con (quote [fp]) _ , _) = R.returnTC nothing + w (xs'@(R.con (quote _fp∷_) tl') , y) = + match2Vargs tl' >>= λ (xs , x) → + R.catchTC + (R.noConstraints $ R.unify + (R.def (quote sym) (x v∷ [])) y + >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) + >>= λ (xs* , xs*↑) → + R.returnTC + (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) + , xs*↑ ) + )) + (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) + ((xs* v∷ y v∷ [])))) + <$> mb-invol n xs') + w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) +mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) + +mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) +mb-invol' = mb-invol magicNumber + + +redList : ℕ → R.Term → R.TC (List R.Term) +redList = h + where + h : ℕ → R.Term → R.TC (List R.Term) + h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) + h (suc k) t = + (mb-invol' t) >>= + Mb.rec + (R.returnTC []) + λ (t' , t'↓) → do + p' ← h k t'↓ + R.returnTC (t' ∷ p') + + +redList' : R.Term → R.TC (List R.Term) +redList' = redList magicNumber + + +pattern fp[_] x = [fp] fp∷ x + +FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ +FP⟦ [fp] ⟧ = refl +FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ + +-- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ +-- FP⟦ [] ⟧t = refl +-- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ +-- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ + + +fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ +fpF→T [fp] = [fp] +fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ + +fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ +fpT→F _ [fp] = [fp] +fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ +fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ +fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ + +fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → + FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ +fpT≡F [fp] = refl +fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x +fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = + hcomp + (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ k ∧ ~ i) + }) + (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j + ; (j = i0) → a₀ + ; (j = i1) → x (~ i ∧ k) + }) + (fpT≡F fp i j)) + +_fp++_ : ∀ {x y z} + → FlatPath A false x y + → FlatPath A false y z + → FlatPath A false x z +x fp++ [fp] = x +x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ + +fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) + (fp' : FlatPath A false a₁ a₂) + → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ +fp++∙ fp [fp] = sym (rUnit _) +fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') + +module PE ([T] : [Typeₙ]) where + + data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where + reflExpr : ∀ {A x} → PathExpr A x x + atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y + compExpr : ∀ {A x y z w} + → PathExpr A x y → PathExpr A y z → PathExpr A z w + → PathExpr A x w + congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') + → PathExpr A x y + → PathExpr A' (f x) (f y) + + + PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → + (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) + PE⟦ reflExpr ⟧ = refl + PE⟦ atomExpr p ⟧ = cong _ p + PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ + PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) + + cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) + → PathExpr A₁ a₀ a₁ + + → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) + (cast↓ [T] A (f a₁)) + cong-flat f reflExpr = [fp] + cong-flat f (atomExpr p) = fp[ cong _ p ] + cong-flat f (compExpr x x₁ x₂) = + cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ + cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x + + flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ + → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) + flat⟦ reflExpr ⟧ = [fp] + flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] + flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ + flat⟦ congExpr f x ⟧ = cong-flat f x + + + cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → + (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → + (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ + FP⟦ cong-flat f pe ⟧ + cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) + cong-flat≡ (atomExpr p) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ + cong-flat≡ (compExpr pe pe₁ pe₂) f = + (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ + cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ + (λ i → doubleCompPath-elim + (cong-flat≡ pe f i) + (cong-flat≡ pe₁ f i) + (cong-flat≡ pe₂ f i) i)) + ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) + (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) + ∙∙ fp++∙ _ (cong-flat f pe₂) + cong-flat≡ (congExpr f₁ pe) f = + cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) + + pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → + PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ + pe≡flat reflExpr = refl + pe≡flat (atomExpr p) = lUnit _ + pe≡flat (compExpr pe pe₁ pe₂) = + (λ i → doubleCompPath-elim + (pe≡flat pe i) + (pe≡flat pe₁ i) + (pe≡flat pe₂ i) i) + ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) + ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ + + pe≡flat (congExpr f pe) = cong-flat≡ pe f + + +module PathTrm (A B : Type ℓ) where + data PathTrm : Type ℓ where + reflTrm : PathTrm + atomTrm : A → PathTrm + compTrm : PathTrm → PathTrm → PathTrm → PathTrm + congTrm : B → PathTrm → PathTrm + + module showPathTrm (showA : A → String) (showB : B → String) where + showPT : PathTrm → String + showPT reflTrm = "refl" + showPT (atomTrm x) = showA x + showPT (compTrm x x₁ x₂) = + "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" + showPT (congTrm x x₁) = + "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" + + +module _ {ℓ ℓ'} + {A B : Type ℓ} + {A' B' : Type ℓ'} + (f : A → R.TC A') + (g : B → R.TC B') where + open PathTrm + mmapPT : PathTrm A B → R.TC $ PathTrm A' B' + mmapPT reflTrm = pure reflTrm + mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ + mmapPT (compTrm x x₁ x₂) = + ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ + mmapPT (congTrm x x₁) = + ⦇ congTrm (g x) (mmapPT x₁) ⦈ + + +module RTrm = PathTrm R.Term R.Term +module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) +module StrTrm = PathTrm String String + +showRTrmTC : RTrm.PathTrm → R.TC String +showRTrmTC = + mmapPT + (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +showRTrmTC' : RTrm'.PathTrm → R.TC String +showRTrmTC' = + let q = λ (k , t) → + R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) + in mmapPT + q q + >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) + +module _ ([T] : [Typeₙ]) where + reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) + reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} + + + atomExpr' : ∀ A {x y} → (p : x ≡ y) → + PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) + + compExpr' : ∀ A {x y z w} + → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w + → PE.PathExpr [T] A x w + compExpr' A = PE.compExpr {[T] = [T]} {A = A} + + congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') + → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) + → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) + (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) + congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} + (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x + +getEnd : ∀ {x y : A} → x ≡ y → A +getEnd p = p i0 + +module tryPathE --([T] : [Typeₙ]) + (TC[T]trm : R.Term) + ([TC[T]trm] [fns] : List R.Term) where + + inferA : R.Term → R.TC ℕ + inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + + inferA' : R.Term → R.TC R.Term + inferA' = inferA >=> R.quoteTC + + + + try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + + + tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + let x₀ = R.def (quote getEnd) v[ t ] + A ← inferA' t + ⦇ (RTrm.reflTrm , R.def (quote reflExpr') + (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ - depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} - → IsPathTrmDeas p → ℕ - depthIsPathTrmDeas nilTrmDeasRefl = zero - depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) + tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + (t1 , t1') ← (try≡ k t1) + (t2 , t2') ← (try≡ k t2) + (t3 , t3') ← (try≡ k t3) + A ← inferA' t + pure (RTrm.compTrm t1 t2 t3 , + R.def (quote compExpr') + (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) - hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool - hasRedexes nilTrmInvolRefl = false - hasRedexes (consTrmInvol x q) = hasRedexes x - hasRedexes (involTrmInvol x q) = true + where - Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ - Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl - Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → - IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q - IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' - IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' - IsPathTrmDeas∙ p' (consTrmDeas q' q'') = - let (_ , u) = IsPathTrmDeas∙ p' q' - in _ , consTrmDeas u q'' + tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t*) ← h tm + A ← inferA' t* + A' ← inferA' t + (t0 , t0') ← (try≡ k t*) + pure (RTrm.congTrm f t0 , + R.def (quote congExpr') + (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) - Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} + where - Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl - Invol→Deas↓ (consTrmInvol x q) = - IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) - Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) - ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl - ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl - ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = - let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r - in _ , isCompTrmReg (isCompTrmReg pR qR) rR , - λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i - ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) - ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl - ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p - ⟦ isCompTrmReg p' q' ⟧da = - let (_ , qD) = ⟦ q' ⟧da - (_ , pD) = ⟦ p' ⟧da - (_ , p∙qD) = IsPathTrmDeas∙ pD qD - in _ , p∙qD + atom : R.Term → R.TC (RTrm.PathTrm × R.Term) + atom x = do + A ← inferA' x + pure (RTrm.atomTrm x , R.def (quote atomExpr') + (TC[T]trm v∷ A v∷ x v∷ [])) - ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) - ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da - module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) - (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where - IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → - (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → - (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) - IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ - IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ - IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = - cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ - sym (∼assoc _ _ _) + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) - ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → - fst ⟦ p' ⟧da ≡ p - ⟦ isReflTrmReg ⟧da≡ = refl - ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) - ⟦ isCompTrmReg p' q' ⟧da≡ = - IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ - daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da - daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) +module tryTermE ([TC[T]trm] [fns] : List R.Term) where + try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) -module _ {A : Type ℓ} where - open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public - open PT≡ rUnit lUnit public + tryRefl : R.Term → R.TC (RTrm'.PathTrm) + tryRefl t = do + _ ← R.noConstraints $ R.checkType + (R.con (quote reflCase) []) + (R.def (quote PathCase) ([ varg t ])) + ⦇ RTrm'.reflTrm ⦈ + tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryComp (suc k) t = do + tm ← R.noConstraints $ R.checkType + (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (t1 , t2 , t3) ← h tm + ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ + where - ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ - ⟦ nilTrmInvolRefl , _ ⟧ti = refl - ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q - ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q - ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti - ⟦ involTrmInvol x q , seg j ⟧ti i = - hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i - ; (i = i1) → q (~ k ∧ ~ j) - ; (i = i0) → ⟦ x , seg j ⟧ti i0 - }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) + h : R.Term → R.TC (R.Term × R.Term × R.Term) + h (R.con (quote compCase) l) = match3Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti - ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti + tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) + tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] + tryCong (suc k) t = + tryAllTC (R.typeError [ "not cong case" ]ₑ) + (zipWithIndex [fns]) + (λ (m , f) → do + tm ← R.noConstraints $ R.checkType + (R.con (quote congCase) (f v∷ [ varg R.unknown ])) + (R.def (quote PathCase) ([ varg t ])) + (_ , t) ← h tm + ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) + where -module _ (A : Type ℓ) (a : A) where - module PTG = PT {A = Unit} (λ _ _ → A) - (a) - (λ _ _ → a) - (λ _ _ _ → a) - (λ _ _ _ → refl) - (λ _ _ _ → refl) + h : R.Term → R.TC (R.Term × R.Term) + h (R.con (quote congCase) l) = match2Vargs l + h _ = R.typeError [ (R.strErr "tryCompFail-h") ] -module PTrm = PTG R.Term R.unknown -module Pℕ = PTG (Bool × ℕ) (true , 0) -module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) + atom : R.Term → R.TC (RTrm'.PathTrm) + atom x = do + k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) + (zipWithIndex [TC[T]trm]) + λ (k , Ty) → + R.checkType x (R.def (quote Path) + (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) + >> pure k + R.returnTC (RTrm'.atomTrm (k , x)) -module _ (f : (Bool × ℕ) → R.Term) where - mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ - mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl - mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) - mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) + try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] + try≡ (suc k) t = + R.catchTC + (tryRefl t) + (R.catchTC (tryComp k t) + (R.catchTC (tryCong k t) (atom t))) -IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) -IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ -ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ +-- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) +-- funTypeView = {!!} -consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ -consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x -consInvolℕ q'@(PT.consTrmDeas x q) x₁ = - decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) - (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) +groupoidSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) +groupoidSolverTerm debugFlag inpF hole = do + -- inp₀ ← wait-for-type inp₀' + -- R.typeError (niceAtomList [) -- >>= wait-for-type + (A' , (t0' , t1')) ← R.inferType hole >>= (get-boundaryWithDom ) >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary" ]) + (λ x → R.returnTC x) -ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl -ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' + -- A ← wait-for-type A' + -- t0 ← wait-for-type t0' + -- t1 ← wait-for-type t1' -module tryPathE where - try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) + let t0 = t0' + let t1 = t1' - tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - R.returnTC (_ , PTrm.isReflTrm) + (AA , _) ← get-boundaryWithDom A' >>= Mb.rec + (R.typeError [ R.strErr "unable to get boundary2" ]) + (λ x → R.returnTC x) + inp₀ ← ([reflectedTy]→reflected[Typeₙ] <$> inferTypesFromFns inpF) >>= wait-for-term - tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (_ , t1') ← try≡ k t1 - (_ , t2') ← try≡ k t2 - (_ , t3') ← try≡ k t3 - R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') + -- (R.typeError [ inp₀ ]ₑ) - where + let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) + -- (R.typeError [ inp ]ₑ) - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) - atom x = R.returnTC (_ , PTrm.isAtomTrm x) + [t] ← reflected[Typeₙ]→[reflectedTy] inp + [f] ← reflected[Fns]→[reflectedFn] inpF + (r0 , r0') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t0) + (r1 , r1') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t1) - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) (atom t)) - -lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y -lamWrap' p i = p i - - -lamWrap : R.Term → R.Term -lamWrap t = R.def (quote lamWrap') (t v∷ []) - -unLam : R.Term → R.TC R.Term -unLam (R.lam _ (R.abs _ t)) = R.returnTC t -unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) - -appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) -appendIfUniqe x [] = R.returnTC [ x ] -appendIfUniqe x xs@(x₁ ∷ xs') = do - x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= - λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) - x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam - sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam - R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ - ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) - ( - R.catchTC - (R.extendContext "i" (varg (R.def (quote I) [])) $ - ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) - (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) - ) - -comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) -comparePathTerms x x₁ = do - x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= - λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) - x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= - λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) - sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam - R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ - ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) - ( - R.catchTC - (R.extendContext "i" (varg (R.def (quote I) [])) $ - ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) - (R.returnTC nothing) - ) - -concatUniq : List R.Term → List R.Term → R.TC (List R.Term) -concatUniq [] = R.returnTC -concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ - -indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) -indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) -indexOfAlways t xs@(x ∷ xs') = - comparePathTerms t x >>= - Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) - (λ b → R.returnTC (xs , (b , 0))) - -unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → - (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) -unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) -unMapAtoms l (PT.isAtomTrm x) = - do (l' , y) ← indexOfAlways x l - R.returnTC (l' , _ , Pℕ.isAtomTrm y) -unMapAtoms l (PT.isCompTrm e e₁ e₂) = - do (l' , _ , e') ← unMapAtoms l e - (l'' , _ , e₁') ← unMapAtoms l' e₁ - (l''' , _ , e₂') ← unMapAtoms l'' e₂ - (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) - - -unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term -unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] -unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) -unquotePathTrm (PT.isCompTrm x x₁ x₂) = - let x' = unquotePathTrm x - x₁' = unquotePathTrm x₁ - x₂' = unquotePathTrm x₂ - in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) - -module _ (l : List R.Term) where - lk : (Bool × ℕ) → R.Term - lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) - where - z = lookupWithDefault R.unknown l n - - - - mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) - mkTrmsInvol' zero _ = [] - mkTrmsInvol' (suc k) u = - let pi = ℕDeas→ℕInvol u - in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] - - mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) - mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x - - unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term - -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) - -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) - unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] - unquoteTrmInvol (PT.consTrmInvol x q) = - R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) - unquoteTrmInvol (PT.involTrmInvol x q) = - R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) - - mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) - mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) - ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x - - foldPathCompTerm : List R.Term → R.Term - foldPathCompTerm [] = R.def (quote refl) [] - foldPathCompTerm (x ∷ []) = x - foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) - - mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) - mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) - - -groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag hole = do - - (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - (r0 , r0') ← tryPathE.try≡ 100 t0 - (r1 , r1') ← tryPathE.try≡ 100 t1 + -- (aL' , (_ , e0)) ← unMapAtoms [] r0' + -- (aL , (_ , e1)) ← unMapAtoms aL' r1' + -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) + -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + show0 ← showRTrmTC r0 + show1 ← showRTrmTC r1 - (aL' , (_ , e0)) ← unMapAtoms [] r0' - (aL , (_ , e1)) ← unMapAtoms aL' r1' - let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) + red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' + red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' - let dbgInfo = (R.strErr "LHS: ") ∷ (R.termErr $ t0) - ∷ (R.strErr "\n") - ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - ∷ (R.strErr "\n") - ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - ∷ (R.strErr "\n") - ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - ∷ (R.strErr "\n") - ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - ∷ (R.strErr "\n") - ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - ∷ (R.strErr "\n") - ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - ∷ (R.strErr "\n") - ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - ∷ (R.strErr "\n") - ∷ (niceAtomList aL) + let invPa0 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red0 + invPa1 = Li.map + (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) + red1 + let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl + ("LHS: ") ∷ₑ show0 ∷nl + ("RHS: ") ∷ₑ show1 ∷nl + (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) + -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) + -- ∷ (R.strErr "\n") + -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - let (_ , iP0) = mkTrmInvol aL e0deas - (eqs1 , iP1) = mkTrmInvol aL e1deas + -- ∷ (R.strErr "\n") + -- (niceAtomList aL) - let dbgInfo = dbgInfo ++ ((R.strErr "\n") ∷ niceEqsList eqs1) + let finalTrm0 = + just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) + ∷ invPa0 + finalTrm1 = + just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) + ∷ invPa1 - let pa0 = R.def (quote _∙_) - (R.def (quote daSingl) ((unquotePathTrm r0') v∷ []) - v∷ iP0 v∷ [] ) - pa1 = R.def (quote _∙_) - (R.def (quote daSingl) ((unquotePathTrm r1') v∷ []) - v∷ iP1 v∷ [] ) - let final = (R.def (quote _∙∙_∙∙_) - (pa0 - v∷ R.def (quote refl) [] - v∷ R.def (quote sym) (pa1 v∷ []) v∷ [] )) + let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms + (finalTrm0 ++ symPathTerms finalTrm1) + pure (finalTrm , dbgInfo) - R.returnTC (final , dbgInfo) - where - - niceEq : ℕ → Pℕ.IsPathTrmInvol _ _ → List R.ErrorPart - niceEq k x = R.strErr (primShowNat k <> " : ") - ∷ R.strErr (PℕS.showIPTI x) - ∷ [ R.strErr "\n" ] - - niceEqsList' : ℕ → List (Pℕ.IsPathTrmInvol _ _) → List R.ErrorPart - niceEqsList' k [] = [] - niceEqsList' k (x ∷ xs) = - niceEq k x ++ niceEqsList' (suc k) xs - - niceEqsList = niceEqsList' 0 - -groupoidSolverMain : Bool → R.Term → R.TC Unit -groupoidSolverMain debugFlag hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type +groupoidSolverMain : Bool → R.Term → R.Term → R.TC Unit +groupoidSolverMain debugFlag inpF' hole = do + inpF ← castTo[Fns] inpF' + inp ← inferTypesFromFns inpF + R.commitTC + ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-term hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag hole' + + (solution , msg) ← groupoidSolverTerm debugFlag inpF hole' R.catchTC (R.unify hole solution) (R.typeError msg) -squareSolverMain : Bool → R.Term → R.TC Unit -squareSolverMain debugFlag hole = do - ty ← R.inferType hole >>= wait-for-type +squareSolverMain : Bool → R.Term → R.Term → R.TC Unit +squareSolverMain debugFlag inpF' hole = do + inpF ← castTo[Fns] inpF' + inp ← inferTypesFromFns inpF + ty ← R.inferType hole >>= wait-for-term hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - - (solution , msg) ← groupoidSolverTerm debugFlag hole' + >>= R.normalise + (solution , msg) ← groupoidSolverTerm debugFlag inpF hole' R.catchTC (R.unify hole' solution) (R.typeError msg) @@ -491,14 +781,16 @@ squareSolverMain debugFlag hole = do R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) macro - solveGroupoidDebug : R.Term → R.TC Unit - solveGroupoidDebug = groupoidSolverMain true - solveSquareDebug : R.Term → R.TC Unit - solveSquareDebug = squareSolverMain false + ≡! : R.Term → R.TC Unit + ≡! = groupoidSolverMain true (R.con (quote [fn]) []) + + ≡!cong : R.Term → R.Term → R.TC Unit + ≡!cong = groupoidSolverMain true + - solveGroupoid : R.Term → R.TC Unit - solveGroupoid = groupoidSolverMain true + sq! : R.Term → R.TC Unit + sq! = squareSolverMain false (R.con (quote [fn]) []) - solveSquare : R.Term → R.TC Unit - solveSquare = squareSolverMain false + sq!cong : R.Term → R.Term → R.TC Unit + sq!cong = squareSolverMain true diff --git a/Cubical/Tactics/PathSolver/Solver2.agda b/Cubical/Tactics/PathSolver/Solver2.agda deleted file mode 100644 index 965d207851..0000000000 --- a/Cubical/Tactics/PathSolver/Solver2.agda +++ /dev/null @@ -1,647 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Solver2 where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - -private - variable - ℓ ℓ' : Level - - -infixr 5 _∷Tω_ - -data [Typeₙ] : Typeω where - [] : [Typeₙ] - _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] - -ℓ[Typeₙ] : [Typeₙ] → Level -ℓ[Typeₙ] [] = ℓ-zero -ℓ[Typeₙ] (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (ℓ[Typeₙ] x₁) - -lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (ℓ[Typeₙ] Ts) -lookupTₙ [] x = Unit -lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {ℓ[Typeₙ] Ts} x₁ -lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) - -data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) - compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) - congCase : ∀ {ℓ ℓ' A A'} (f : A → A') {x y} - → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - - --- module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} --- (p : Path {ℓ} A x y) where - --- pathCaseCongTest : PathCase (cong f p) --- pathCaseCongTest = congCase f λ z → p z - -data PathExprω : {ℓ : Level} (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflExpr : ∀ {ℓ A x} → PathExprω {ℓ} A (refl {x = x}) - atomExpr : ∀ {ℓ A x y} → (p : x ≡ y) → PathExprω {ℓ} A p - compExpr : ∀ {ℓ A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} - → PathExprω {ℓ} A p → PathExprω {ℓ} A q → PathExprω {ℓ} A r - → PathExprω {ℓ} A (p ∙∙ q ∙∙ r) - congExpr : ∀ {ℓ ℓ' A A' x y} → (f : A → A') {p : x ≡ y} - → PathExprω {ℓ} A p - → PathExprω {ℓ'} A' (cong f p) - - -data PathExpr {ℓ : Level} : (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Type (ℓ-suc ℓ) where - reflExpr : ∀ {A x} → PathExpr A (refl {x = x}) - atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A p - compExpr : ∀ {A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} - → PathExpr A p → PathExpr A q → PathExpr A r - → PathExpr A (p ∙∙ q ∙∙ r) - congExpr : ∀ {A A' x y} → (f : A → A') {p : x ≡ y} - → PathExpr {ℓ} A {x} {y} p - → PathExpr A' (cong f p) - - - -PEω⟦_⟧ : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → - PathExprω A p → a₀ ≡ a₁ -PEω⟦ reflExpr ⟧ = refl -PEω⟦ atomExpr p ⟧ = p -PEω⟦ compExpr x x₁ x₂ ⟧ = PEω⟦ x ⟧ ∙∙ PEω⟦ x₁ ⟧ ∙∙ PEω⟦ x₂ ⟧ -PEω⟦ congExpr f x ⟧ = cong f PEω⟦ x ⟧ - - -PE⟦_⟧ : {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → - PathExpr A p → a₀ ≡ a₁ -PE⟦ reflExpr ⟧ = refl -PE⟦ atomExpr p ⟧ = p -PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ -PE⟦ congExpr f x ⟧ = cong f PE⟦ x ⟧ - - -record ΣPathExprω : Typeω where - field - ℓa : Level - A : Type ℓa - a₀ a₁ : A - p : a₀ ≡ a₁ - expr : PathExprω A p - - ΣPEω⟦_⟧ : a₀ ≡ a₁ - ΣPEω⟦_⟧ = PEω⟦ expr ⟧ - -record ΣPathExpr {ℓ} : Type (ℓ-suc ℓ) where - constructor Σpe - field - {A} : Type ℓ - {a₀ a₁} : A - {p} : a₀ ≡ a₁ - expr : PathExpr A p - - ΣPE⟦_⟧ : a₀ ≡ a₁ - ΣPE⟦_⟧ = PE⟦ expr ⟧ - - - -module PathTrm (A B : Type ℓ) where - data PathTrm : Type ℓ where - reflTrm : PathTrm - atomTrm : A → PathTrm - compTrm : PathTrm → PathTrm → PathTrm → PathTrm - congTrm : B → PathTrm → PathTrm - - module showPathTrm (showA : A → String) (showB : B → String) where - showPT : PathTrm → String - showPT reflTrm = "refl" - showPT (atomTrm x) = showA x - showPT (compTrm x x₁ x₂) = - "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" - showPT (congTrm x x₁) = - "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" - - -module _ {ℓ ℓ'} - {A B : Type ℓ} - {A' B' : Type ℓ'} - (f : A → R.TC A') - (g : B → R.TC B') where - open PathTrm - mmapPT : PathTrm A B → R.TC $ PathTrm A' B' - mmapPT reflTrm = pure reflTrm - mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ - mmapPT (compTrm x x₁ x₂) = - ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ - mmapPT (congTrm x x₁) = - ⦇ congTrm (g x) (mmapPT x₁) ⦈ - - -module RTrm = PathTrm R.Term R.Term -module StrTrm = PathTrm String String - --- unqouteRTrm : ∀ {ℓ} → RTrm.PathTrm → R.TC (ΣPathExpr {ℓ}) --- unqouteRTrm PathTrm.reflTrm = pure $ Σpe {A = Unit*} reflExpr --- unqouteRTrm (PathTrm.atomTrm x) = --- ⦇ (Σpe ∘ atomExpr) {!x!} ⦈ --- unqouteRTrm (PathTrm.compTrm x x₁ x₂) = {!!} --- unqouteRTrm (PathTrm.congTrm x x₁) = {!!} - -showRTrmTC : RTrm.PathTrm → R.TC String -showRTrmTC = - mmapPT - (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -module tryPathE where - - try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm) - - - tryRefl : R.Term → R.TC (RTrm.PathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - R.returnTC (RTrm.reflTrm) - - tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (t1') ← try≡ k t1 - (t2') ← try≡ k t2 - (t3') ← try≡ k t3 - R.returnTC (RTrm.compTrm t1' t2' t3') - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - atom : R.Term → R.TC (RTrm.PathTrm) - atom x = R.returnTC (RTrm.atomTrm x) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) (atom t)) - - - - - - - - - - -- data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - -- isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) - -- isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p - -- isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} - -- → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) - - -- data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - -- nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) - -- consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) - - -- data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where - -- nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x - -- consTrmInvol : ∀ {x y z : _} → - -- IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z - -- involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → - -- IsPathTrmInvol x y - - --- module show (showA : ∀ {ℓ} A {a₀ a₁ : A} → (p : a₀ ≡ a₁) → String) where --- showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String --- showIPT isReflTrm = "refl" --- showIPT (isAtomTrm x) = showA x --- showIPT (isCompTrm isReflTrm x₁ x₂) = --- "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" --- showIPT (isCompTrm x x₁ isReflTrm) = --- "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" --- showIPT (isCompTrm x x₁ x₂) = --- "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" - --- -- showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String - --- -- showIPTD nilTrmDeasRefl = "refl" --- -- showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q - --- -- showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String --- -- showIPTI nilTrmInvolRefl = "refl" --- -- showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q --- -- showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" - - --- -- depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} --- -- → IsPathTrmDeas p → ℕ --- -- depthIsPathTrmDeas nilTrmDeasRefl = zero --- -- depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) - --- -- hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool --- -- hasRedexes nilTrmInvolRefl = false --- -- hasRedexes (consTrmInvol x q) = hasRedexes x --- -- hasRedexes (involTrmInvol x q) = true - --- -- Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ --- -- Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl --- -- Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q - --- -- IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → --- -- IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q --- -- IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' --- -- IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' --- -- IsPathTrmDeas∙ p' (consTrmDeas q' q'') = --- -- let (_ , u) = IsPathTrmDeas∙ p' q' --- -- in _ , consTrmDeas u q'' - - --- -- Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} - --- -- Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl --- -- Invol→Deas↓ (consTrmInvol x q) = --- -- IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) --- -- Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x - --- -- ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) --- -- ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl --- -- ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl --- -- ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = --- -- let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r --- -- in _ , isCompTrmReg (isCompTrmReg pR qR) rR , --- -- λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i - - --- -- ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) --- -- ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl --- -- ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p --- -- ⟦ isCompTrmReg p' q' ⟧da = --- -- let (_ , qD) = ⟦ q' ⟧da --- -- (_ , pD) = ⟦ p' ⟧da --- -- (_ , p∙qD) = IsPathTrmDeas∙ pD qD --- -- in _ , p∙qD - --- -- ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) --- -- ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da --- -- module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) --- -- (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where - --- -- IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → --- -- (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → --- -- (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) --- -- IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ --- -- IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ --- -- IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = --- -- cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ --- -- sym (∼assoc _ _ _) - --- -- ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → --- -- fst ⟦ p' ⟧da ≡ p --- -- ⟦ isReflTrmReg ⟧da≡ = refl --- -- ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) --- -- ⟦ isCompTrmReg p' q' ⟧da≡ = --- -- IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ - --- -- daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da --- -- daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) - - - --- -- module _ {A : Type ℓ} where - --- -- open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public --- -- open PT≡ rUnit lUnit public - - - --- -- ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ --- -- ⟦ nilTrmInvolRefl , _ ⟧ti = refl --- -- ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q --- -- ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q --- -- ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti --- -- ⟦ involTrmInvol x q , seg j ⟧ti i = --- -- hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i --- -- ; (i = i1) → q (~ k ∧ ~ j) --- -- ; (i = i0) → ⟦ x , seg j ⟧ti i0 --- -- }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) - --- -- ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti --- -- ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti - - - --- module _ (A : Type ℓ) (a : A) where --- module PTG = PT {A = Unit} (λ _ _ → A) --- (a) --- (λ _ _ → a) --- (λ _ _ _ → a) --- (λ _ _ _ → refl) --- (λ _ _ _ → refl) - --- module PTrm = PTG R.Term R.unknown - --- module Pℕ = PTG (Bool × ℕ) (true , 0) - --- module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) - - --- -- module _ (f : (Bool × ℕ) → R.Term) where --- -- mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ --- -- mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl --- -- mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) --- -- mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) - - --- IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) --- IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ - --- -- ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ - --- -- consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ --- -- consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x --- -- consInvolℕ q'@(PT.consTrmDeas x q) x₁ = --- -- decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) --- -- (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) - - - --- -- ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl --- -- ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' - --- module tryPathE where - --- try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) - - --- tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- tryRefl t = do --- _ ← R.noConstraints $ R.checkType --- (R.con (quote reflCase) []) --- (R.def (quote PathCase) ([ varg t ])) --- R.returnTC (_ , PTrm.isReflTrm) - --- tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] --- tryComp (suc k) t = do --- tm ← R.noConstraints $ R.checkType --- (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) --- (R.def (quote PathCase) ([ varg t ])) --- (t1 , t2 , t3) ← h tm --- (_ , t1') ← try≡ k t1 --- (_ , t2') ← try≡ k t2 --- (_ , t3') ← try≡ k t3 --- R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') - --- where - --- h : R.Term → R.TC (R.Term × R.Term × R.Term) --- h (R.con (quote compCase) l) = match3Vargs l --- h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - --- atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- atom x = R.returnTC (_ , PTrm.isAtomTrm x) - - --- try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] --- try≡ (suc k) t = --- R.catchTC --- (tryRefl t) --- (R.catchTC (tryComp k t) (atom t)) - --- lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y --- lamWrap' p i = p i - - --- lamWrap : R.Term → R.Term --- lamWrap t = R.def (quote lamWrap') (t v∷ []) - --- unLam : R.Term → R.TC R.Term --- unLam (R.lam _ (R.abs _ t)) = R.returnTC t --- unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) - --- appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) --- appendIfUniqe x [] = R.returnTC [ x ] --- appendIfUniqe x xs@(x₁ ∷ xs') = do --- x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) --- x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) --- ( --- R.catchTC --- (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) --- (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) --- ) - --- comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) --- comparePathTerms x x₁ = do --- x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) --- x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) --- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) --- ( --- R.catchTC --- (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) --- (R.returnTC nothing) --- ) - --- concatUniq : List R.Term → List R.Term → R.TC (List R.Term) --- concatUniq [] = R.returnTC --- concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ - --- indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) --- indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) --- indexOfAlways t xs@(x ∷ xs') = --- comparePathTerms t x >>= --- Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) --- (λ b → R.returnTC (xs , (b , 0))) - --- unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → --- (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) --- unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) --- unMapAtoms l (PT.isAtomTrm x) = --- do (l' , y) ← indexOfAlways x l --- R.returnTC (l' , _ , Pℕ.isAtomTrm y) --- unMapAtoms l (PT.isCompTrm e e₁ e₂) = --- do (l' , _ , e') ← unMapAtoms l e --- (l'' , _ , e₁') ← unMapAtoms l' e₁ --- (l''' , _ , e₂') ← unMapAtoms l'' e₂ --- (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) - - --- -- unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term --- -- unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] --- -- unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) --- -- unquotePathTrm (PT.isCompTrm x x₁ x₂) = --- -- let x' = unquotePathTrm x --- -- x₁' = unquotePathTrm x₁ --- -- x₂' = unquotePathTrm x₂ --- -- in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) - --- -- module _ (l : List R.Term) where --- -- lk : (Bool × ℕ) → R.Term --- -- lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) --- -- where --- -- z = lookupWithDefault R.unknown l n - - - --- -- mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) --- -- mkTrmsInvol' zero _ = [] --- -- mkTrmsInvol' (suc k) u = --- -- let pi = ℕDeas→ℕInvol u --- -- in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] - --- -- mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) --- -- mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x - --- -- unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term --- -- -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) --- -- -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) --- -- unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] --- -- unquoteTrmInvol (PT.consTrmInvol x q) = --- -- R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) --- -- unquoteTrmInvol (PT.involTrmInvol x q) = --- -- R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) - --- -- mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) --- -- mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) --- -- ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x - --- -- foldPathCompTerm : List R.Term → R.Term --- -- foldPathCompTerm [] = R.def (quote refl) [] --- -- foldPathCompTerm (x ∷ []) = x --- -- foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) - --- -- mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) --- -- mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) - - -groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag hole = do - - (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - r0 ← tryPathE.try≡ 100 t0 - r1 ← tryPathE.try≡ 100 t1 - - - -- (aL' , (_ , e0)) ← unMapAtoms [] r0' - -- (aL , (_ , e1)) ← unMapAtoms aL' r1' - -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) - show0 ← showRTrmTC r0 - show1 ← showRTrmTC r1 - - let dbgInfo = ("LHS: ") ∷ₑ show0 ∷nl - ("RHS: ") ∷ₑ show1 ∷nl [] - -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - - -- ∷ (R.strErr "\n") - -- (niceAtomList aL) - - R.typeError dbgInfo - - -groupoidSolverMain : Bool → R.Term → R.TC Unit -groupoidSolverMain debugFlag hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - -squareSolverMain : Bool → R.Term → R.TC Unit -squareSolverMain debugFlag hole = do - ty ← R.inferType hole >>= wait-for-type - hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - - (solution , msg) ← groupoidSolverTerm debugFlag hole' - R.catchTC - (R.unify hole' solution) - (R.typeError msg) - - R.catchTC - (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) - (R.typeError [ R.strErr "xxx" ] ) - where - extractMeta : R.Term → R.TC R.Term - extractMeta (R.def _ tl) = matchVarg tl - extractMeta t = - R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) - -macro - solveGroupoidDebug : R.Term → R.TC Unit - solveGroupoidDebug = groupoidSolverMain true - - -- solveSquareDebug : R.Term → R.TC Unit - -- solveSquareDebug = squareSolverMain false - - -- solveGroupoid : R.Term → R.TC Unit - -- solveGroupoid = groupoidSolverMain true - - -- solveSquare : R.Term → R.TC Unit - -- solveSquare = squareSolverMain false - - --- module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - --- pA pB pC : x ≡ w --- pA = (p ∙∙ refl ∙∙ q) ∙ sym r --- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl --- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - --- pA=pB : pA ≡ pB --- pA=pB = solveGroupoidDebug diff --git a/Cubical/Tactics/PathSolver/Solver3.agda b/Cubical/Tactics/PathSolver/Solver3.agda deleted file mode 100644 index f213d278ea..0000000000 --- a/Cubical/Tactics/PathSolver/Solver3.agda +++ /dev/null @@ -1,806 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Solver3 where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - -private - variable - ℓ ℓ' : Level - A A' : Type ℓ - - -infixr 5 _∷Tω_ - -data [Typeₙ] : Typeω where - [] : [Typeₙ] - _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] - -maxℓ : [Typeₙ] → Level -maxℓ [] = ℓ-zero -maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) - -lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) -lookupTₙ [] x = ⊥* -lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ -lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) - -ℓAt : (Ts : [Typeₙ]) → ℕ → Level -ℓAt [] x = ℓ-zero -ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ -ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x - -TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) -TyAt [] k = ⊥* -TyAt (x ∷Tω Ts) zero = x -TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k - -cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k -cast↓ [] k () -cast↓ (x₁ ∷Tω Ts) zero x = lower x -cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) - -cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y -cast↓Inj {[]} {A = A} {()} -cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift -cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = - cong lift (cast↓Inj {[T] = [T]} {A = A} p) - -cast↓Inj-sec : ∀ {[T] A x y} p → - cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p -cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl -cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = - cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) - -cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → - cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) - ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) - - -cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ -cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = - cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ - - -cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k -cast↑ [] k () -cast↑ (x₁ ∷Tω Ts) zero x = lift x -cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x - --- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) --- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t - - --- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) --- Ts-arr' [] x x₁ = Unit --- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) --- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} --- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} --- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = --- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) - --- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → --- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t --- Ts-arr' Ts s t x x₁ = {!!} - - - -data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) - compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) - congCase : ∀ {ℓ ℓ' A A'} (f : A → A') {x y} - → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - - -module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} - (p : Path {ℓ} A x y) where - - pathCaseCongTest : PathCase (cong f p) - pathCaseCongTest = congCase f λ z → p z - - -infixl 5 _fp∷_ -infixl 5 _fp++_ - -data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where - [] : ∀ {x b} → FlatPath A b x x - _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z - _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y - - -FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ -FP⟦ [] ⟧ = refl -FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ - -pattern fp[_] x = [] fp∷ x - -_fp++_ : ∀ {x y z} - → FlatPath A false x y - → FlatPath A false y z - → FlatPath A false x z -x fp++ [] = x -x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ - -fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) - (fp' : FlatPath A false a₁ a₂) - → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ -fp++∙ fp [] = sym (rUnit _) -fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') - -module _ ([T] : [Typeₙ]) where - - data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where - reflExpr : ∀ {A x} → PathExpr A x x - atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y - compExpr : ∀ {A x y z w} - → PathExpr A x y → PathExpr A y z → PathExpr A z w - → PathExpr A x w - congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') - → PathExpr A x y - → PathExpr A' (f x) (f y) - - - PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → - (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) - PE⟦ reflExpr ⟧ = refl - PE⟦ atomExpr p ⟧ = cong _ p - PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ - PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) - - cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) - → PathExpr A₁ a₀ a₁ - - → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) - (cast↓ [T] A (f a₁)) - cong-flat f reflExpr = [] - cong-flat f (atomExpr p) = fp[ cong _ p ] - cong-flat f (compExpr x x₁ x₂) = - cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ - cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x - - flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ - → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) - flat⟦ reflExpr ⟧ = [] - flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] - flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ - flat⟦ congExpr f x ⟧ = cong-flat f x - - - cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → - (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → - (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ - FP⟦ cong-flat f pe ⟧ - cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) - cong-flat≡ (atomExpr p) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ - cong-flat≡ (compExpr pe pe₁ pe₂) f = - (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ - cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ - (λ i → doubleCompPath-elim - (cong-flat≡ pe f i) - (cong-flat≡ pe₁ f i) - (cong-flat≡ pe₂ f i) i)) - ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) - (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) - ∙∙ fp++∙ _ (cong-flat f pe₂) - cong-flat≡ (congExpr f₁ pe) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) - - pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → - PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ - pe≡flat reflExpr = refl - pe≡flat (atomExpr p) = lUnit _ - pe≡flat (compExpr pe pe₁ pe₂) = - (λ i → doubleCompPath-elim - (pe≡flat pe i) - (pe≡flat pe₁ i) - (pe≡flat pe₂ i) i) - ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) - ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ - - pe≡flat (congExpr f pe) = cong-flat≡ pe f - --- data PathExprω : {ℓ : Level} (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where --- reflExpr : ∀ {ℓ A x} → PathExprω {ℓ} A (refl {x = x}) --- atomExpr : ∀ {ℓ A x y} → (p : x ≡ y) → PathExprω {ℓ} A p --- compExpr : ∀ {ℓ A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} --- → PathExprω {ℓ} A p → PathExprω {ℓ} A q → PathExprω {ℓ} A r --- → PathExprω {ℓ} A (p ∙∙ q ∙∙ r) --- congExpr : ∀ {ℓ ℓ' A A' x y} → (f : A → A') {p : x ≡ y} --- → PathExprω {ℓ} A p --- → PathExprω {ℓ'} A' (cong f p) - - --- data PathExpr {ℓ : Level} : (A : Type ℓ) {a₀ a₁ : A} → a₀ ≡ a₁ → Type (ℓ-suc ℓ) where --- reflExpr : ∀ {A x} → PathExpr A (refl {x = x}) --- atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A p --- compExpr : ∀ {A x y z w} → {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} --- → PathExpr A p → PathExpr A q → PathExpr A r --- → PathExpr A (p ∙∙ q ∙∙ r) --- congExpr : ∀ {A A' x y} → (f : A → A') {p : x ≡ y} --- → PathExpr {ℓ} A {x} {y} p --- → PathExpr A' (cong f p) - - - --- PEω⟦_⟧ : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → {p : a₀ ≡ a₁} → --- PathExprω A p → a₀ ≡ a₁ --- PEω⟦ reflExpr ⟧ = refl --- PEω⟦ atomExpr p ⟧ = p --- PEω⟦ compExpr x x₁ x₂ ⟧ = PEω⟦ x ⟧ ∙∙ PEω⟦ x₁ ⟧ ∙∙ PEω⟦ x₂ ⟧ --- PEω⟦ congExpr f x ⟧ = cong f PEω⟦ x ⟧ - - - --- record ΣPathExprω : Typeω where --- field --- ℓa : Level --- A : Type ℓa --- a₀ a₁ : A --- p : a₀ ≡ a₁ --- expr : PathExprω A p - --- ΣPEω⟦_⟧ : a₀ ≡ a₁ --- ΣPEω⟦_⟧ = PEω⟦ expr ⟧ - --- record ΣPathExpr {ℓ} : Type (ℓ-suc ℓ) where --- constructor Σpe --- field --- {A} : Type ℓ --- {a₀ a₁} : A --- {p} : a₀ ≡ a₁ --- expr : PathExpr A p - --- ΣPE⟦_⟧ : a₀ ≡ a₁ --- ΣPE⟦_⟧ = PE⟦ expr ⟧ - - - -module PathTrm (A B : Type ℓ) where - data PathTrm : Type ℓ where - reflTrm : PathTrm - atomTrm : A → PathTrm - compTrm : PathTrm → PathTrm → PathTrm → PathTrm - congTrm : B → PathTrm → PathTrm - - module showPathTrm (showA : A → String) (showB : B → String) where - showPT : PathTrm → String - showPT reflTrm = "refl" - showPT (atomTrm x) = showA x - showPT (compTrm x x₁ x₂) = - "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" - showPT (congTrm x x₁) = - "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" - - -module _ {ℓ ℓ'} - {A B : Type ℓ} - {A' B' : Type ℓ'} - (f : A → R.TC A') - (g : B → R.TC B') where - open PathTrm - mmapPT : PathTrm A B → R.TC $ PathTrm A' B' - mmapPT reflTrm = pure reflTrm - mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ - mmapPT (compTrm x x₁ x₂) = - ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ - mmapPT (congTrm x x₁) = - ⦇ congTrm (g x) (mmapPT x₁) ⦈ - - -module RTrm = PathTrm R.Term R.Term -module StrTrm = PathTrm String String - --- unqouteRTrm : ∀ {ℓ} → RTrm.PathTrm → R.TC (ΣPathExpr {ℓ}) --- unqouteRTrm PathTrm.reflTrm = pure $ Σpe {A = Unit*} reflExpr --- unqouteRTrm (PathTrm.atomTrm x) = --- ⦇ (Σpe ∘ atomExpr) {!x!} ⦈ --- unqouteRTrm (PathTrm.compTrm x x₁ x₂) = {!!} --- unqouteRTrm (PathTrm.congTrm x x₁) = {!!} - -showRTrmTC : RTrm.PathTrm → R.TC String -showRTrmTC = - mmapPT - (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - - - -module tryPathE where - - try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm) - - - tryRefl : R.Term → R.TC (RTrm.PathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - R.returnTC (RTrm.reflTrm) - - tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (t1') ← try≡ k t1 - (t2') ← try≡ k t2 - (t3') ← try≡ k t3 - R.returnTC (RTrm.compTrm t1' t2' t3') - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - atom : R.Term → R.TC (RTrm.PathTrm) - atom x = R.returnTC (RTrm.atomTrm x) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) (atom t)) - - - - - - - - - - -- data IsPathTrmReg : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - -- isReflTrmReg : ∀ {x} → IsPathTrmReg (∼refl {x = x}) - -- isAtomTrmReg : ∀ {x y} → (p : x ∼ y) → IsPathTrmReg p - -- isCompTrmReg : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} - -- → IsPathTrmReg p → IsPathTrmReg q → IsPathTrmReg (p ∼∙ q) - - -- data IsPathTrmDeas : {a₀ a₁ : A} → a₀ ∼ a₁ → Type (ℓ-max ℓ ℓ') where - -- nilTrmDeasRefl : ∀ {x} → IsPathTrmDeas (∼refl {x = x}) - -- consTrmDeas : ∀ {x y z : _} → {p : x ∼ y} → IsPathTrmDeas p → (q : y ∼ z) → IsPathTrmDeas (p ∼∙ q) - - -- data IsPathTrmInvol : (a₀ a₁ : A) → Type (ℓ-max ℓ ℓ') where - -- nilTrmInvolRefl : ∀ {x} → IsPathTrmInvol x x - -- consTrmInvol : ∀ {x y z : _} → - -- IsPathTrmInvol x y → (q : y ∼ z) → IsPathTrmInvol x z - -- involTrmInvol : ∀ {x y z : _} → IsPathTrmInvol x y → (q : y ∼ z) → - -- IsPathTrmInvol x y - - --- module show (showA : ∀ {ℓ} A {a₀ a₁ : A} → (p : a₀ ≡ a₁) → String) where --- showIPT : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → String --- showIPT isReflTrm = "refl" --- showIPT (isAtomTrm x) = showA x --- showIPT (isCompTrm isReflTrm x₁ x₂) = --- "(" <> showIPT x₁ <> "∙" <> showIPT x₂ <> ")" --- showIPT (isCompTrm x x₁ isReflTrm) = --- "(" <> showIPT x <> "∙'" <> showIPT x₁ <> ")" --- showIPT (isCompTrm x x₁ x₂) = --- "(" <> showIPT x <> "∙∙" <> showIPT x₁ <> "∙∙" <> showIPT x₂ <> ")" - --- -- showIPTD : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmDeas p → String - --- -- showIPTD nilTrmDeasRefl = "refl" --- -- showIPTD (consTrmDeas x q) = showIPTD x <> "∙" <> showA q - --- -- showIPTI : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → String --- -- showIPTI nilTrmInvolRefl = "refl" --- -- showIPTI (consTrmInvol x q) = showIPTI x <> "∙" <> showA q --- -- showIPTI (involTrmInvol x q) = showIPTI x <> "∙⟦" <> showA q <> "∙" <> showA q <> "⁻¹⟧" - - --- -- depthIsPathTrmDeas : ∀ {a₀ a₁ : A} → ∀ {p : a₀ ∼ a₁} --- -- → IsPathTrmDeas p → ℕ --- -- depthIsPathTrmDeas nilTrmDeasRefl = zero --- -- depthIsPathTrmDeas (consTrmDeas x q) = suc (depthIsPathTrmDeas x) - --- -- hasRedexes : ∀ {a₀ a₁} → IsPathTrmInvol a₀ a₁ → Bool --- -- hasRedexes nilTrmInvolRefl = false --- -- hasRedexes (consTrmInvol x q) = hasRedexes x --- -- hasRedexes (involTrmInvol x q) = true - --- -- Deas→Invol : ∀ {a₀ a₁ : A} → ∀ {p} → IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} p → IsPathTrmInvol a₀ a₁ --- -- Deas→Invol nilTrmDeasRefl = nilTrmInvolRefl --- -- Deas→Invol (consTrmDeas x q) = consTrmInvol (Deas→Invol x) q - --- -- IsPathTrmDeas∙ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → --- -- IsPathTrmDeas p → IsPathTrmDeas q → Σ _ λ p∙q → IsPathTrmDeas {x} {z} p∙q --- -- IsPathTrmDeas∙ p' nilTrmDeasRefl = _ , p' --- -- IsPathTrmDeas∙ nilTrmDeasRefl q'@(consTrmDeas _ _) = _ , q' --- -- IsPathTrmDeas∙ p' (consTrmDeas q' q'') = --- -- let (_ , u) = IsPathTrmDeas∙ p' q' --- -- in _ , consTrmDeas u q'' - - --- -- Invol→Deas↓ : ∀ {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Σ _ $ IsPathTrmDeas {a₀ = a₀} {a₁ = a₁} - --- -- Invol→Deas↓ nilTrmInvolRefl = _ , nilTrmDeasRefl --- -- Invol→Deas↓ (consTrmInvol x q) = --- -- IsPathTrmDeas∙ (snd (Invol→Deas↓ x)) (consTrmDeas nilTrmDeasRefl q) --- -- Invol→Deas↓ (involTrmInvol x q) = Invol→Deas↓ x - --- -- ⟦_⟧r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ λ r → (IsPathTrmReg r × (p ≡ r))) --- -- ⟦ isReflTrm ⟧r = ∼refl , isReflTrmReg , refl --- -- ⟦ isAtomTrm p ⟧r = p , isAtomTrmReg _ , refl --- -- ⟦ isCompTrm {p = p} {q = q} {r = r} p' q' r' ⟧r = --- -- let (_ , pR , p=) = ⟦ p' ⟧r ; (_ , qR , q=) = ⟦ q' ⟧r ; (_ , rR , r=) = ⟦ r' ⟧r --- -- in _ , isCompTrmReg (isCompTrmReg pR qR) rR , --- -- λ i → ∼doubleCompPath-elim (p= i) (q= i) (r= i) i - - --- -- ⟦_⟧da : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrmReg p → (Σ _ λ r → (IsPathTrmDeas r)) --- -- ⟦ isReflTrmReg ⟧da = _ , nilTrmDeasRefl --- -- ⟦ isAtomTrmReg p ⟧da = _ , consTrmDeas nilTrmDeasRefl p --- -- ⟦ isCompTrmReg p' q' ⟧da = --- -- let (_ , qD) = ⟦ q' ⟧da --- -- (_ , pD) = ⟦ p' ⟧da --- -- (_ , p∙qD) = IsPathTrmDeas∙ pD qD --- -- in _ , p∙qD - --- -- ⟦_⟧da∘r : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → IsPathTrm p → (Σ _ IsPathTrmDeas) --- -- ⟦ x ⟧da∘r = ⟦ fst (snd (⟦ x ⟧r)) ⟧da --- -- module PT≡ (∼rUnit : ∀ {x y} → (p : x ∼ y) → p ≡ p ∼∙ ∼refl) --- -- (∼lUnit : ∀ {x y} → (p : x ∼ y) → p ≡ ∼refl ∼∙ p) where - --- -- IsPathTrmDeas∙≡ : ∀ {x y z : _} → {p : x ∼ y} {q : y ∼ z} → --- -- (p' : IsPathTrmDeas p) → (q' : IsPathTrmDeas q) → --- -- (fst (IsPathTrmDeas∙ p' q') ≡ (p ∼∙ q)) --- -- IsPathTrmDeas∙≡ _ nilTrmDeasRefl = ∼rUnit _ --- -- IsPathTrmDeas∙≡ nilTrmDeasRefl (consTrmDeas p q) = ∼lUnit _ --- -- IsPathTrmDeas∙≡ (consTrmDeas p q) (consTrmDeas p' q') = --- -- cong (_∼∙ q') ( (IsPathTrmDeas∙≡ (consTrmDeas p q) p')) ∙ --- -- sym (∼assoc _ _ _) - --- -- ⟦_⟧da≡ : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (p' : IsPathTrmReg p) → --- -- fst ⟦ p' ⟧da ≡ p --- -- ⟦ isReflTrmReg ⟧da≡ = refl --- -- ⟦ isAtomTrmReg _ ⟧da≡ = sym (∼lUnit _) --- -- ⟦ isCompTrmReg p' q' ⟧da≡ = --- -- IsPathTrmDeas∙≡ (snd ⟦ p' ⟧da) (snd ⟦ q' ⟧da) ∙ cong₂ _∼∙_ ⟦ p' ⟧da≡ ⟦ q' ⟧da≡ - --- -- daSingl : {a₀ a₁ : A} → {p : a₀ ∼ a₁} → (q : IsPathTrm p) → p ≡ fst ⟦ fst (snd ⟦ q ⟧r) ⟧da --- -- daSingl x = let (_ , x' , x=) = ⟦ x ⟧r in x= ∙ sym (⟦ x' ⟧da≡) - - - --- -- module _ {A : Type ℓ} where - --- -- open PT {A = A} _≡_ refl _∙_ _∙∙_∙∙_ doubleCompPath-elim assoc public --- -- open PT≡ rUnit lUnit public - - - --- -- ⟦_,_⟧ti : {a₀ a₁ : A} → IsPathTrmInvol a₀ a₁ → Interval → a₀ ≡ a₁ --- -- ⟦ nilTrmInvolRefl , _ ⟧ti = refl --- -- ⟦ consTrmInvol x q , 𝓲 ⟧ti = ⟦ x , 𝓲 ⟧ti ∙ q --- -- ⟦ involTrmInvol x q , zero ⟧ti = (⟦ x , zero ⟧ti ∙ q) ∙ sym q --- -- ⟦ involTrmInvol x q , one ⟧ti = ⟦ x , one ⟧ti --- -- ⟦ involTrmInvol x q , seg j ⟧ti i = --- -- hcomp (λ k → λ { (j = i1) → ⟦ x , one ⟧ti i --- -- ; (i = i1) → q (~ k ∧ ~ j) --- -- ; (i = i0) → ⟦ x , seg j ⟧ti i0 --- -- }) (compPath-filler ⟦ x , seg j ⟧ti q (~ j) i) - --- -- ⟦_⟧ti≡ : {a₀ a₁ : A} → (x : IsPathTrmInvol a₀ a₁) → ⟦ x , zero ⟧ti ≡ ⟦ x , one ⟧ti --- -- ⟦_⟧ti≡ x i = ⟦ x , (seg i) ⟧ti - - - --- module _ (A : Type ℓ) (a : A) where --- module PTG = PT {A = Unit} (λ _ _ → A) --- (a) --- (λ _ _ → a) --- (λ _ _ _ → a) --- (λ _ _ _ → refl) --- (λ _ _ _ → refl) - --- module PTrm = PTG R.Term R.unknown - --- module Pℕ = PTG (Bool × ℕ) (true , 0) - --- module PℕS = Pℕ.show (λ (b , i) → let v = mkNiceVar i in if b then v else (v <> "⁻¹")) - - --- -- module _ (f : (Bool × ℕ) → R.Term) where --- -- mapPTG : Pℕ.IsPathTrmInvol _ _ → PTrm.IsPathTrmInvol _ _ --- -- mapPTG PT.nilTrmInvolRefl = PT.nilTrmInvolRefl --- -- mapPTG (PT.consTrmInvol x q) = PT.consTrmInvol (mapPTG x) (f q) --- -- mapPTG (PT.involTrmInvol x q) = PT.involTrmInvol (mapPTG x) (f q) - - --- IsRedex? : ∀ x x' → Dec (Path (Bool × ℕ) x (not (fst x') , snd x')) --- IsRedex? _ _ = discreteΣ 𝟚._≟_ (λ _ → discreteℕ) _ _ - --- -- ℕDeas→ℕInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → Pℕ.IsPathTrmInvol _ _ - --- -- consInvolℕ : ∀ {p} → Pℕ.IsPathTrmDeas p → (Bool × ℕ) → Pℕ.IsPathTrmInvol _ _ --- -- consInvolℕ PT.nilTrmDeasRefl x = PT.consTrmInvol PT.nilTrmInvolRefl x --- -- consInvolℕ q'@(PT.consTrmDeas x q) x₁ = --- -- decRec (λ _ → Pℕ.involTrmInvol (ℕDeas→ℕInvol x) q) --- -- (λ _ → Pℕ.consTrmInvol (ℕDeas→ℕInvol q') x₁) (IsRedex? q x₁ ) - - - --- -- ℕDeas→ℕInvol PT.nilTrmDeasRefl = PT.nilTrmInvolRefl --- -- ℕDeas→ℕInvol (PT.consTrmDeas p' q') = consInvolℕ p' q' - --- module tryPathE where - --- try≡ : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) - - --- tryRefl : R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- tryRefl t = do --- _ ← R.noConstraints $ R.checkType --- (R.con (quote reflCase) []) --- (R.def (quote PathCase) ([ varg t ])) --- R.returnTC (_ , PTrm.isReflTrm) - --- tryComp : ℕ → R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] --- tryComp (suc k) t = do --- tm ← R.noConstraints $ R.checkType --- (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) --- (R.def (quote PathCase) ([ varg t ])) --- (t1 , t2 , t3) ← h tm --- (_ , t1') ← try≡ k t1 --- (_ , t2') ← try≡ k t2 --- (_ , t3') ← try≡ k t3 --- R.returnTC (_ , PTrm.isCompTrm t1' t2' t3') - --- where - --- h : R.Term → R.TC (R.Term × R.Term × R.Term) --- h (R.con (quote compCase) l) = match3Vargs l --- h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - --- atom : R.Term → R.TC (Σ _ PTrm.IsPathTrm) --- atom x = R.returnTC (_ , PTrm.isAtomTrm x) - - --- try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] --- try≡ (suc k) t = --- R.catchTC --- (tryRefl t) --- (R.catchTC (tryComp k t) (atom t)) - --- lamWrap' : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ y → x ≡ y --- lamWrap' p i = p i - - --- lamWrap : R.Term → R.Term --- lamWrap t = R.def (quote lamWrap') (t v∷ []) - --- unLam : R.Term → R.TC R.Term --- unLam (R.lam _ (R.abs _ t)) = R.returnTC t --- unLam t = R.typeError ((R.strErr "unLam-fail") ∷ [ R.termErr t ]) - --- appendIfUniqe : R.Term → List R.Term → R.TC (List R.Term) --- appendIfUniqe x [] = R.returnTC [ x ] --- appendIfUniqe x xs@(x₁ ∷ xs') = do --- x' ← R.checkType x (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "aiu x'" ]) --- x₁' ← R.checkType x₁ (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC xs)) --- ( --- R.catchTC --- (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC xs)) --- (appendIfUniqe x xs' >>= (R.returnTC ∘ (x₁ ∷_)) ) --- ) - --- comparePathTerms : R.Term → R.Term → R.TC (Maybe Bool) --- comparePathTerms x x₁ = do --- x' ← R.withNormalisation true $ R.checkType (lamWrap x) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError [ R.strErr "cpt x'" ]) --- x₁' ← R.withNormalisation true $ R.checkType (lamWrap x₁) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= --- λ u → R.catchTC (unLam u) (R.typeError (R.strErr "cpt x₁'" ∷ R.termErr x ∷ [ R.termErr x₁ ])) --- sym[x₁'] ← R.checkType (R.def (quote sym) [ varg x₁ ]) (R.def (quote _≡_) (R.unknown v∷ R.unknown v∷ [])) >>= unLam --- R.catchTC (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') (x₁') >> R.returnTC (just true))) --- ( --- R.catchTC --- (R.extendContext "i" (varg (R.def (quote I) [])) $ --- ( R.noConstraints $ R.unify (x') sym[x₁'] >> R.returnTC (just false))) --- (R.returnTC nothing) --- ) - --- concatUniq : List R.Term → List R.Term → R.TC (List R.Term) --- concatUniq [] = R.returnTC --- concatUniq (x ∷ x₂) x₁ = appendIfUniqe x x₁ >>= concatUniq x₂ - --- indexOfAlways : R.Term → List R.Term → R.TC ((List R.Term) × (Bool × ℕ)) --- indexOfAlways t [] = R.returnTC ([ t ] , (true , 0)) --- indexOfAlways t xs@(x ∷ xs') = --- comparePathTerms t x >>= --- Mb.rec ((λ (l , (b , k) ) → (x ∷ l) , (b , (suc k))) <$> indexOfAlways t xs' ) --- (λ b → R.returnTC (xs , (b , 0))) - --- unMapAtoms : List R.Term → ∀ {p} → PTrm.IsPathTrm p → --- (R.TC ((List R.Term) × (Σ _ Pℕ.IsPathTrm))) --- unMapAtoms l PT.isReflTrm = R.returnTC (l , _ , Pℕ.isReflTrm) --- unMapAtoms l (PT.isAtomTrm x) = --- do (l' , y) ← indexOfAlways x l --- R.returnTC (l' , _ , Pℕ.isAtomTrm y) --- unMapAtoms l (PT.isCompTrm e e₁ e₂) = --- do (l' , _ , e') ← unMapAtoms l e --- (l'' , _ , e₁') ← unMapAtoms l' e₁ --- (l''' , _ , e₂') ← unMapAtoms l'' e₂ --- (R.returnTC (l''' , _ , Pℕ.isCompTrm e' e₁' e₂')) - - --- -- unquotePathTrm : ∀ {p} → PTrm.IsPathTrm p → R.Term --- -- unquotePathTrm PT.isReflTrm = R.con (quote (isReflTrm)) [] --- -- unquotePathTrm (PT.isAtomTrm p) = R.con (quote isAtomTrm) (p v∷ []) --- -- unquotePathTrm (PT.isCompTrm x x₁ x₂) = --- -- let x' = unquotePathTrm x --- -- x₁' = unquotePathTrm x₁ --- -- x₂' = unquotePathTrm x₂ --- -- in R.con (quote isCompTrm) (x' v∷ x₁' v∷ x₂' v∷ []) - --- -- module _ (l : List R.Term) where --- -- lk : (Bool × ℕ) → R.Term --- -- lk (b , n) = if b then z else (R.def (quote sym) (z v∷ [])) --- -- where --- -- z = lookupWithDefault R.unknown l n - - - --- -- mkTrmsInvol' : ∀ {p} → ℕ → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) --- -- mkTrmsInvol' zero _ = [] --- -- mkTrmsInvol' (suc k) u = --- -- let pi = ℕDeas→ℕInvol u --- -- in if (Pℕ.hasRedexes pi) then (pi ∷ mkTrmsInvol' k (snd (Pℕ.Invol→Deas↓ pi))) else [] - --- -- mkTrmsInvol* : ∀ {p} → Pℕ.IsPathTrmDeas p → List (Pℕ.IsPathTrmInvol _ _) --- -- mkTrmsInvol* x = mkTrmsInvol' (Pℕ.depthIsPathTrmDeas x) x - --- -- unquoteTrmInvol : PTrm.IsPathTrmInvol tt tt → R.Term --- -- -- unquoteTrmInvol (PT.nilTrmInvol p) = R.con (quote nilTrmInvol) (p v∷ []) --- -- -- unquoteTrmInvol (PT.nilInvolTrmInvol p) = R.con (quote nilInvolTrmInvol) (p v∷ []) --- -- unquoteTrmInvol PT.nilTrmInvolRefl = R.con (quote nilTrmInvolRefl) [] --- -- unquoteTrmInvol (PT.consTrmInvol x q) = --- -- R.con (quote consTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) --- -- unquoteTrmInvol (PT.involTrmInvol x q) = --- -- R.con (quote involTrmInvol) (unquoteTrmInvol x v∷ q v∷ []) - --- -- mkTrmsInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → List (R.Term) --- -- mkTrmsInvol x = Li.map ((λ y → R.def (quote ⟦_⟧ti≡) (y v∷ [])) --- -- ∘ unquoteTrmInvol ∘ mapPTG lk) $ mkTrmsInvol* x - --- -- foldPathCompTerm : List R.Term → R.Term --- -- foldPathCompTerm [] = R.def (quote refl) [] --- -- foldPathCompTerm (x ∷ []) = x --- -- foldPathCompTerm (x ∷ xs@(_ ∷ _)) = R.def (quote _∙_) (x v∷ foldPathCompTerm xs v∷ []) - --- -- mkTrmInvol : ∀ {p} → Pℕ.IsPathTrmDeas p → (List (Pℕ.IsPathTrmInvol _ _) × R.Term) --- -- mkTrmInvol x = ( mkTrmsInvol* x) , foldPathCompTerm (mkTrmsInvol x) - - -groupoidSolverTerm : Bool → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag hole = do - - (t0 , t1) ← R.inferType hole >>= wait-for-type >>= (get-boundary ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - r0 ← tryPathE.try≡ 100 t0 - r1 ← tryPathE.try≡ 100 t1 - - - -- (aL' , (_ , e0)) ← unMapAtoms [] r0' - -- (aL , (_ , e1)) ← unMapAtoms aL' r1' - -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) - show0 ← showRTrmTC r0 - show1 ← showRTrmTC r1 - - let dbgInfo = ("LHS: ") ∷ₑ show0 ∷nl - ("RHS: ") ∷ₑ show1 ∷nl [] - -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - - -- ∷ (R.strErr "\n") - -- (niceAtomList aL) - - R.typeError dbgInfo - - -groupoidSolverMain : Bool → R.Term → R.TC Unit -groupoidSolverMain debugFlag hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - -squareSolverMain : Bool → R.Term → R.TC Unit -squareSolverMain debugFlag hole = do - ty ← R.inferType hole >>= wait-for-type - hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - - (solution , msg) ← groupoidSolverTerm debugFlag hole' - R.catchTC - (R.unify hole' solution) - (R.typeError msg) - - R.catchTC - (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) - (R.typeError [ R.strErr "xxx" ] ) - where - extractMeta : R.Term → R.TC R.Term - extractMeta (R.def _ tl) = matchVarg tl - extractMeta t = - R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) - -macro - solveGroupoidDebug : R.Term → R.TC Unit - solveGroupoidDebug = groupoidSolverMain true - - -- solveSquareDebug : R.Term → R.TC Unit - -- solveSquareDebug = squareSolverMain false - - -- solveGroupoid : R.Term → R.TC Unit - -- solveGroupoid = groupoidSolverMain true - - -- solveSquare : R.Term → R.TC Unit - -- solveSquare = squareSolverMain false - - --- module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - --- pA pB pC : x ≡ w --- pA = (p ∙∙ refl ∙∙ q) ∙ sym r --- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl --- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - --- pA=pB : pA ≡ pB --- pA=pB = solveGroupoidDebug diff --git a/Cubical/Tactics/PathSolver/Solver4.agda b/Cubical/Tactics/PathSolver/Solver4.agda deleted file mode 100644 index b5ad972dde..0000000000 --- a/Cubical/Tactics/PathSolver/Solver4.agda +++ /dev/null @@ -1,877 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Solver4 where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - -private - variable - ℓ ℓ' : Level - A A' : Type ℓ - - -infixr 5 _∷Tω_ - -infixr 5 _∷fn_ - - -data [Typeₙ] : Typeω where - [Tω] : [Typeₙ] - _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] - -data [Fns] : Typeω where - [fn] : [Fns] - _∷fn_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] - - -reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) -reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] _ = - R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ - - --- macro --- test-refl[T]macro : R.Term → R.Term → R.TC Unit --- test-refl[T]macro inp hole = do --- l ← reflected[Typeₙ]→[reflectedTy] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[T] : Unit --- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) - - -reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) -reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ - --- macro --- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit --- test-refl[Fn]macro inp hole = do --- l ← reflected[Fns]→[reflectedFn] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[Fn] : Unit --- test-refl[Fn] = test-refl[Fn]macro --- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) - - -maxℓ : [Typeₙ] → Level -maxℓ [Tω] = ℓ-zero -maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) - -lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) -lookupTₙ [Tω] x = ⊥* -lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ -lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) - -ℓAt : (Ts : [Typeₙ]) → ℕ → Level -ℓAt [Tω] x = ℓ-zero -ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ -ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x - -TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) -TyAt [Tω] k = ⊥* -TyAt (x ∷Tω Ts) zero = x -TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k - -cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k -cast↓ [Tω] k () -cast↓ (x₁ ∷Tω Ts) zero x = lower x -cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) - -cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y -cast↓Inj {[Tω]} {A = A} {()} -cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift -cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = - cong lift (cast↓Inj {[T] = [T]} {A = A} p) - -cast↓Inj-sec : ∀ {[T] A x y} p → - cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p -cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl -cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = - cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) - -cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → - cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) - ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) - - -cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ -cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = - cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ - - -cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k -cast↑ [Tω] k () -cast↑ (x₁ ∷Tω Ts) zero x = lift x -cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x - --- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) --- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t - - --- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) --- Ts-arr' [] x x₁ = Unit --- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) --- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} --- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} --- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = --- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) - --- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → --- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t --- Ts-arr' Ts s t x x₁ = {!!} - - - -data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) - compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) - congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') - → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - - -module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} - (p : Path {ℓ} A x y) where - - -- pathCaseCongTest : PathCase λ i → f (p i) - -- pathCaseCongTest = congCase f {!!} - - -infixl 5 _fp∷_ -infixl 5 _fp++_ - -data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where - [fp] : ∀ {x b} → FlatPath A b x x - _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z - _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y - -magicNumber = 100 - -mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) -mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing -mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w - where - w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) - w (R.con (quote [fp]) _ , _) = R.returnTC nothing - w (xs'@(R.con (quote _fp∷_) tl') , y) = - match2Vargs tl' >>= λ (xs , x) → - R.catchTC - (R.noConstraints $ R.unify - (R.def (quote sym) (x v∷ [])) y - >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) - >>= λ (xs* , xs*↑) → - R.returnTC - (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) - , xs*↑ ) - )) - (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) - ((xs* v∷ y v∷ [])))) - <$> mb-invol n xs') - w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) -mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) - -mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol' = mb-invol magicNumber - - -redList : ℕ → R.Term → R.TC (List R.Term) -redList = h - where - h : ℕ → R.Term → R.TC (List R.Term) - h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) - h (suc k) t = - (mb-invol' t) >>= - Mb.rec - (R.returnTC []) - λ (t' , t'↓) → do - p' ← h k t'↓ - R.returnTC (t' ∷ p') - - -redList' : R.Term → R.TC (List R.Term) -redList' = redList magicNumber - - -pattern fp[_] x = [fp] fp∷ x - -FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ -FP⟦ [fp] ⟧ = refl -FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ - --- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ --- FP⟦ [] ⟧t = refl --- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ --- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ - - -fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ -fpF→T [fp] = [fp] -fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ - -fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ -fpT→F _ [fp] = [fp] -fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ -fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ -fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ - -fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → - FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ -fpT≡F [fp] = refl -fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x -fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = - hcomp - (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ k ∧ ~ i) - }) - (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ i ∧ k) - }) - (fpT≡F fp i j)) - -_fp++_ : ∀ {x y z} - → FlatPath A false x y - → FlatPath A false y z - → FlatPath A false x z -x fp++ [fp] = x -x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ - -fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) - (fp' : FlatPath A false a₁ a₂) - → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ -fp++∙ fp [fp] = sym (rUnit _) -fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') - -module PE ([T] : [Typeₙ]) where - - data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where - reflExpr : ∀ {A x} → PathExpr A x x - atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y - compExpr : ∀ {A x y z w} - → PathExpr A x y → PathExpr A y z → PathExpr A z w - → PathExpr A x w - congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') - → PathExpr A x y - → PathExpr A' (f x) (f y) - - - PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → - (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) - PE⟦ reflExpr ⟧ = refl - PE⟦ atomExpr p ⟧ = cong _ p - PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ - PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) - - cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) - → PathExpr A₁ a₀ a₁ - - → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) - (cast↓ [T] A (f a₁)) - cong-flat f reflExpr = [fp] - cong-flat f (atomExpr p) = fp[ cong _ p ] - cong-flat f (compExpr x x₁ x₂) = - cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ - cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x - - flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ - → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) - flat⟦ reflExpr ⟧ = [fp] - flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] - flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ - flat⟦ congExpr f x ⟧ = cong-flat f x - - - cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → - (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → - (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ - FP⟦ cong-flat f pe ⟧ - cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) - cong-flat≡ (atomExpr p) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ - cong-flat≡ (compExpr pe pe₁ pe₂) f = - (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ - cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ - (λ i → doubleCompPath-elim - (cong-flat≡ pe f i) - (cong-flat≡ pe₁ f i) - (cong-flat≡ pe₂ f i) i)) - ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) - (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) - ∙∙ fp++∙ _ (cong-flat f pe₂) - cong-flat≡ (congExpr f₁ pe) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) - - pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → - PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ - pe≡flat reflExpr = refl - pe≡flat (atomExpr p) = lUnit _ - pe≡flat (compExpr pe pe₁ pe₂) = - (λ i → doubleCompPath-elim - (pe≡flat pe i) - (pe≡flat pe₁ i) - (pe≡flat pe₂ i) i) - ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) - ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ - - pe≡flat (congExpr f pe) = cong-flat≡ pe f - - -module PathTrm (A B : Type ℓ) where - data PathTrm : Type ℓ where - reflTrm : PathTrm - atomTrm : A → PathTrm - compTrm : PathTrm → PathTrm → PathTrm → PathTrm - congTrm : B → PathTrm → PathTrm - - module showPathTrm (showA : A → String) (showB : B → String) where - showPT : PathTrm → String - showPT reflTrm = "refl" - showPT (atomTrm x) = showA x - showPT (compTrm x x₁ x₂) = - "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" - showPT (congTrm x x₁) = - "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" - - -module _ {ℓ ℓ'} - {A B : Type ℓ} - {A' B' : Type ℓ'} - (f : A → R.TC A') - (g : B → R.TC B') where - open PathTrm - mmapPT : PathTrm A B → R.TC $ PathTrm A' B' - mmapPT reflTrm = pure reflTrm - mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ - mmapPT (compTrm x x₁ x₂) = - ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ - mmapPT (congTrm x x₁) = - ⦇ congTrm (g x) (mmapPT x₁) ⦈ - - -module RTrm = PathTrm R.Term R.Term -module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) -module StrTrm = PathTrm String String - -showRTrmTC : RTrm.PathTrm → R.TC String -showRTrmTC = - mmapPT - (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -showRTrmTC' : RTrm'.PathTrm → R.TC String -showRTrmTC' = - let q = λ (k , t) → - R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) - in mmapPT - q q - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -module _ ([T] : [Typeₙ]) where - reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) - reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} - - - atomExpr' : ∀ A {x y} → (p : x ≡ y) → - PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) - - compExpr' : ∀ A {x y z w} - → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w - → PE.PathExpr [T] A x w - compExpr' A = PE.compExpr {[T] = [T]} {A = A} - - congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') - → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) - (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) - congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} - (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x - -getEnd : ∀ {x y : A} → x ≡ y → A -getEnd p = p i0 - -module tryPathE --([T] : [Typeₙ]) - (TC[T]trm : R.Term) - ([TC[T]trm] [fns] : List R.Term) where - - inferA : R.Term → R.TC ℕ - inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - - inferA' : R.Term → R.TC R.Term - inferA' = inferA >=> R.quoteTC - - - - try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - - - tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - let x₀ = R.def (quote getEnd) v[ t ] - A ← inferA' t - ⦇ (RTrm.reflTrm , R.def (quote reflExpr') - (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (t1 , t1') ← (try≡ k t1) - (t2 , t2') ← (try≡ k t2) - (t3 , t3') ← (try≡ k t3) - A ← inferA' t - pure (RTrm.compTrm t1 t2 t3 , - R.def (quote compExpr') - (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t*) ← h tm - A ← inferA' t* - A' ← inferA' t - (t0 , t0') ← (try≡ k t*) - pure (RTrm.congTrm f t0 , - R.def (quote congExpr') - (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm.PathTrm × R.Term) - atom x = do - A ← inferA' x - pure (RTrm.atomTrm x , R.def (quote atomExpr') - (TC[T]trm v∷ A v∷ x v∷ [])) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - -module tryTermE --([T] : [Typeₙ]) - --(TC[T]trm : R.Term) - ([TC[T]trm] [fns] : List R.Term) where - - try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) - - - tryRefl : R.Term → R.TC (RTrm'.PathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - ⦇ RTrm'.reflTrm ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t) ← h tm - ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm'.PathTrm) - atom x = do - k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - R.returnTC (RTrm'.atomTrm (k , x)) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - --- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) --- funTypeView = {!!} - -groupoidSolverTerm : Bool → R.Term → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag inp₀ inpF hole = do - - -- inp₀ ← wait-for-type inp₀' - - -- R.typeError (niceAtomList [) - (A' , (t0' , t1')) ← R.inferType hole >>= wait-for-type >>= (get-boundaryWithDom ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - -- A ← wait-for-type A' - -- t0 ← wait-for-type t0' - -- t1 ← wait-for-type t1' - - let t0 = t0' - let t1 = t1' - - (AA , _) ← get-boundaryWithDom A' >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary2" ]) - (λ x → R.returnTC x) - - let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) - -- (R.typeError [ inp ]ₑ) - - - - [t] ← reflected[Typeₙ]→[reflectedTy] inp - [f] ← reflected[Fns]→[reflectedFn] inpF - - (r0 , r0') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t0) - (r1 , r1') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t1) - - - - -- (aL' , (_ , e0)) ← unMapAtoms [] r0' - -- (aL , (_ , e1)) ← unMapAtoms aL' r1' - -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) - show0 ← showRTrmTC r0 - show1 ← showRTrmTC r1 - - red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' - red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' - - - let invPa0 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red0 - invPa1 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red1 - - let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl - ("LHS: ") ∷ₑ show0 ∷nl - ("RHS: ") ∷ₑ show1 ∷nl - (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) - -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - - -- ∷ (R.strErr "\n") - -- (niceAtomList aL) - - let finalTrm0 = - just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) - ∷ invPa0 - - finalTrm1 = - just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) - ∷ invPa1 - - let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms - (finalTrm0 ++ symPathTerms finalTrm1) - - pure (finalTrm , dbgInfo) - - -groupoidSolverMain : Bool → R.Term → R.Term → R.Term → R.TC Unit -groupoidSolverMain debugFlag inp inpF hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag inp inpF hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - -groupoidSolverMain' : Bool → R.Term → R.TC Unit -groupoidSolverMain' debugFlag hole = do - - let inpF = R.con (quote [fn]) [] - let inp = R.con (quote [Tω]) [] - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - hole' ← R.withNormalisation true $ R.checkType hole ty - R.commitTC - (solution , msg) ← groupoidSolverTerm debugFlag inp inpF hole' - -- R.typeError [ solution ]ₑ - R.catchTC - (R.unify hole solution) - (R.typeError msg) - - --- -- squareSolverMain : Bool → R.Term → R.Term → R.TC Unit --- -- squareSolverMain debugFlag inp hole = do --- -- ty ← R.inferType hole >>= wait-for-type --- -- hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - --- -- (solution , msg) ← groupoidSolverTerm debugFlag inp hole' --- -- R.catchTC --- -- (R.unify hole' solution) --- -- (R.typeError msg) - --- -- R.catchTC --- -- (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) --- -- (R.typeError [ R.strErr "xxx" ] ) --- -- where --- -- extractMeta : R.Term → R.TC R.Term --- -- extractMeta (R.def _ tl) = matchVarg tl --- -- extractMeta t = --- -- R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) - -macro - solveGroupoidDebug : R.Term → R.Term → R.Term → R.TC Unit - solveGroupoidDebug = groupoidSolverMain true - - ≡! : R.Term → R.TC Unit - ≡! = groupoidSolverMain' true - --- -- solveSquareDebug : R.Term → R.TC Unit --- -- solveSquareDebug = squareSolverMain false - --- -- solveGroupoid : R.Term → R.TC Unit --- -- solveGroupoid = groupoidSolverMain true - --- -- solveSquare : R.Term → R.TC Unit --- -- solveSquare = squareSolverMain false - - -module SimpleTest where - - module _ (n : ℕ) where - - data AA : Type where - x y z w : AA - p p' : x ≡ y - q : y ≡ z - q' : z ≡ y - r : z ≡ w - - - - pA pB : Path (AA 3) (x) (w) - pA = ((refl ∙ p) ∙ q) ∙ r - pB = ((refl ∙ p) ∙ q) ∙ r - - - pA=pB : Path (Path (AA 3) x y) - (refl ∙ p) - (p) - pA=pB = ≡! - - -module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - - pA pB pC : x ≡ w - pA = (p ∙∙ refl ∙∙ q) ∙ sym r - pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - - pA=pB : pA ≡ pB - pA=pB = ≡! - - -module Examples' (A : ℕ → Type ℓ) - (x y z w : ∀ {n} → A n) - (p p' : ∀ {n} → x {n} ≡ y) - (q : ∀ {n} → y {n} ≡ z) - (q' : ∀ {n} → z {n} ≡ y) - (r : ∀ {n} → w {n} ≡ z) where - - pA pB : Path (A 3) x w - pA = (p ∙∙ refl ∙∙ q) ∙ sym r - pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - - pA=pB : pA ≡ pB - pA=pB = ≡! - - -module Examplesℕ' where - - data AA : (n : ℕ) → Type where - x y z w : ∀ {n} → AA n - p p' : ∀ {n} → x {n} ≡ y - q : ∀ {n} → y {n} ≡ z - q' : ∀ {n} → z {n} ≡ y - r : ∀ {n} → w {n} ≡ z - - - - pA pB : Path (AA 3) (x) (w) - pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) - pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) - ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl - - pA=pB : pA ≡ pB - pA=pB = ≡! - - -module Examplesℕ'' where - - module _ (n : ℕ) where - - data AA : Type where - x y z w : AA - p p' : x ≡ y - q : y ≡ z - q' : z ≡ y - r : w ≡ z - - - - pA pB : Path (AA 3) (x) (w) - pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) - pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) - ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl - - - pA=pB : pA ≡ pB - pA=pB = ≡! - - --- module Examplesℕ where - --- data AA (n : ℕ) : Type where --- x y z w : AA n --- p p' : x ≡ y --- q : y ≡ z --- q' : z ≡ y --- r : w ≡ z - - - --- pA pB : Path (AA 3) (x) (w) --- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) --- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) --- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl --- -- -- -- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - --- pA=pB : pA ≡ pB --- pA=pB = ≡! - - - - --- -- -- -- -- -- -- -- module Examples2 (A B : Type ℓ) (x y z : B) (w : A) (f g : B → A) --- -- -- -- -- -- -- -- (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ f z) where - --- -- -- -- -- -- -- -- pA pB : f x ≡ w --- -- -- -- -- -- -- -- pA = cong f (p ∙∙ refl ∙∙ q) ∙ sym r --- -- -- -- -- -- -- -- pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl - - --- -- -- -- -- -- -- -- pA=pB : pA ≡ pB --- -- -- -- -- -- -- -- pA=pB = solveGroupoidDebug (B ∷Tω A ∷Tω [Tω]) (g ∷fn f ∷fn [fn]) - diff --git a/Cubical/Tactics/PathSolver/Solver5.agda b/Cubical/Tactics/PathSolver/Solver5.agda deleted file mode 100644 index db236cde76..0000000000 --- a/Cubical/Tactics/PathSolver/Solver5.agda +++ /dev/null @@ -1,876 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Solver5 where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - -private - variable - ℓ ℓ' : Level - A A' : Type ℓ - - -infixr 5 _∷Tω_ - -infixr 5 _∷fn_ - - -data [Typeₙ] : Typeω where - [Tω] : [Typeₙ] - _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] - -data [Fns] : Typeω where - [fn] : [Fns] - _∷fn_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] - - -reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) -reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] _ = - R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ - - --- macro --- test-refl[T]macro : R.Term → R.Term → R.TC Unit --- test-refl[T]macro inp hole = do --- l ← reflected[Typeₙ]→[reflectedTy] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[T] : Unit --- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) - - -reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) -reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote _∷fn_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ - --- macro --- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit --- test-refl[Fn]macro inp hole = do --- l ← reflected[Fns]→[reflectedFn] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[Fn] : Unit --- test-refl[Fn] = test-refl[Fn]macro --- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) - - -maxℓ : [Typeₙ] → Level -maxℓ [Tω] = ℓ-zero -maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) - -lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) -lookupTₙ [Tω] x = ⊥* -lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ -lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) - -ℓAt : (Ts : [Typeₙ]) → ℕ → Level -ℓAt [Tω] x = ℓ-zero -ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ -ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x - -TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) -TyAt [Tω] k = ⊥* -TyAt (x ∷Tω Ts) zero = x -TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k - -cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k -cast↓ [Tω] k () -cast↓ (x₁ ∷Tω Ts) zero x = lower x -cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) - -cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y -cast↓Inj {[Tω]} {A = A} {()} -cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift -cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = - cong lift (cast↓Inj {[T] = [T]} {A = A} p) - -cast↓Inj-sec : ∀ {[T] A x y} p → - cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p -cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl -cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = - cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) - -cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → - cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) - ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) - - -cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ -cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = - cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ - - -cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k -cast↑ [Tω] k () -cast↑ (x₁ ∷Tω Ts) zero x = lift x -cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x - --- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) --- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t - - --- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) --- Ts-arr' [] x x₁ = Unit --- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) --- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} --- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} --- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = --- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) - --- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → --- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t --- Ts-arr' Ts s t x x₁ = {!!} - - - -data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) - compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) - congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') - → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - - -module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} - (p : Path {ℓ} A x y) where - - -- pathCaseCongTest : PathCase λ i → f (p i) - -- pathCaseCongTest = congCase f {!!} - - -infixl 5 _fp∷_ -infixl 5 _fp++_ - -data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where - [fp] : ∀ {x b} → FlatPath A b x x - _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z - _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y - -magicNumber = 100 - -mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) -mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing -mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w - where - w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) - w (R.con (quote [fp]) _ , _) = R.returnTC nothing - w (xs'@(R.con (quote _fp∷_) tl') , y) = - match2Vargs tl' >>= λ (xs , x) → - R.catchTC - (R.noConstraints $ R.unify - (R.def (quote sym) (x v∷ [])) y - >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) - >>= λ (xs* , xs*↑) → - R.returnTC - (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) - , xs*↑ ) - )) - (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) - ((xs* v∷ y v∷ [])))) - <$> mb-invol n xs') - w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) -mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) - -mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol' = mb-invol magicNumber - - -redList : ℕ → R.Term → R.TC (List R.Term) -redList = h - where - h : ℕ → R.Term → R.TC (List R.Term) - h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) - h (suc k) t = - (mb-invol' t) >>= - Mb.rec - (R.returnTC []) - λ (t' , t'↓) → do - p' ← h k t'↓ - R.returnTC (t' ∷ p') - - -redList' : R.Term → R.TC (List R.Term) -redList' = redList magicNumber - - -pattern fp[_] x = [fp] fp∷ x - -FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ -FP⟦ [fp] ⟧ = refl -FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ - --- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ --- FP⟦ [] ⟧t = refl --- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ --- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ - - -fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ -fpF→T [fp] = [fp] -fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ - -fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ -fpT→F _ [fp] = [fp] -fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ -fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ -fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ - -fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → - FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ -fpT≡F [fp] = refl -fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x -fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = - hcomp - (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ k ∧ ~ i) - }) - (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ i ∧ k) - }) - (fpT≡F fp i j)) - -_fp++_ : ∀ {x y z} - → FlatPath A false x y - → FlatPath A false y z - → FlatPath A false x z -x fp++ [fp] = x -x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ - -fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) - (fp' : FlatPath A false a₁ a₂) - → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ -fp++∙ fp [fp] = sym (rUnit _) -fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') - -module PE ([T] : [Typeₙ]) where - - data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where - reflExpr : ∀ {A x} → PathExpr A x x - atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y - compExpr : ∀ {A x y z w} - → PathExpr A x y → PathExpr A y z → PathExpr A z w - → PathExpr A x w - congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') - → PathExpr A x y - → PathExpr A' (f x) (f y) - - - PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → - (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) - PE⟦ reflExpr ⟧ = refl - PE⟦ atomExpr p ⟧ = cong _ p - PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ - PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) - - cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) - → PathExpr A₁ a₀ a₁ - - → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) - (cast↓ [T] A (f a₁)) - cong-flat f reflExpr = [fp] - cong-flat f (atomExpr p) = fp[ cong _ p ] - cong-flat f (compExpr x x₁ x₂) = - cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ - cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x - - flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ - → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) - flat⟦ reflExpr ⟧ = [fp] - flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] - flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ - flat⟦ congExpr f x ⟧ = cong-flat f x - - - cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → - (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → - (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ - FP⟦ cong-flat f pe ⟧ - cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) - cong-flat≡ (atomExpr p) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ - cong-flat≡ (compExpr pe pe₁ pe₂) f = - (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ - cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ - (λ i → doubleCompPath-elim - (cong-flat≡ pe f i) - (cong-flat≡ pe₁ f i) - (cong-flat≡ pe₂ f i) i)) - ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) - (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) - ∙∙ fp++∙ _ (cong-flat f pe₂) - cong-flat≡ (congExpr f₁ pe) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) - - pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → - PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ - pe≡flat reflExpr = refl - pe≡flat (atomExpr p) = lUnit _ - pe≡flat (compExpr pe pe₁ pe₂) = - (λ i → doubleCompPath-elim - (pe≡flat pe i) - (pe≡flat pe₁ i) - (pe≡flat pe₂ i) i) - ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) - ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ - - pe≡flat (congExpr f pe) = cong-flat≡ pe f - - -module PathTrm (A B : Type ℓ) where - data PathTrm : Type ℓ where - reflTrm : PathTrm - atomTrm : A → PathTrm - compTrm : PathTrm → PathTrm → PathTrm → PathTrm - congTrm : B → PathTrm → PathTrm - - module showPathTrm (showA : A → String) (showB : B → String) where - showPT : PathTrm → String - showPT reflTrm = "refl" - showPT (atomTrm x) = showA x - showPT (compTrm x x₁ x₂) = - "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" - showPT (congTrm x x₁) = - "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" - - -module _ {ℓ ℓ'} - {A B : Type ℓ} - {A' B' : Type ℓ'} - (f : A → R.TC A') - (g : B → R.TC B') where - open PathTrm - mmapPT : PathTrm A B → R.TC $ PathTrm A' B' - mmapPT reflTrm = pure reflTrm - mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ - mmapPT (compTrm x x₁ x₂) = - ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ - mmapPT (congTrm x x₁) = - ⦇ congTrm (g x) (mmapPT x₁) ⦈ - - -module RTrm = PathTrm R.Term R.Term -module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) -module StrTrm = PathTrm String String - -showRTrmTC : RTrm.PathTrm → R.TC String -showRTrmTC = - mmapPT - (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -showRTrmTC' : RTrm'.PathTrm → R.TC String -showRTrmTC' = - let q = λ (k , t) → - R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) - in mmapPT - q q - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -module _ ([T] : [Typeₙ]) where - reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) - reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} - - - atomExpr' : ∀ A {x y} → (p : x ≡ y) → - PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) - - compExpr' : ∀ A {x y z w} - → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w - → PE.PathExpr [T] A x w - compExpr' A = PE.compExpr {[T] = [T]} {A = A} - - congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') - → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) - (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) - congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} - (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x - -getEnd : ∀ {x y : A} → x ≡ y → A -getEnd p = p i0 - -module tryPathE --([T] : [Typeₙ]) - (TC[T]trm : R.Term) - ([TC[T]trm] [fns] : List R.Term) where - - inferA : R.Term → R.TC ℕ - inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - - inferA' : R.Term → R.TC R.Term - inferA' = inferA >=> R.quoteTC - - - - try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - - - tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - let x₀ = R.def (quote getEnd) v[ t ] - A ← inferA' t - ⦇ (RTrm.reflTrm , R.def (quote reflExpr') - (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (t1 , t1') ← (try≡ k t1) - (t2 , t2') ← (try≡ k t2) - (t3 , t3') ← (try≡ k t3) - A ← inferA' t - pure (RTrm.compTrm t1 t2 t3 , - R.def (quote compExpr') - (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t*) ← h tm - A ← inferA' t* - A' ← inferA' t - (t0 , t0') ← (try≡ k t*) - pure (RTrm.congTrm f t0 , - R.def (quote congExpr') - (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm.PathTrm × R.Term) - atom x = do - A ← inferA' x - pure (RTrm.atomTrm x , R.def (quote atomExpr') - (TC[T]trm v∷ A v∷ x v∷ [])) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - -module tryTermE --([T] : [Typeₙ]) - --(TC[T]trm : R.Term) - ([TC[T]trm] [fns] : List R.Term) where - - try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) - - - tryRefl : R.Term → R.TC (RTrm'.PathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - ⦇ RTrm'.reflTrm ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t) ← h tm - ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm'.PathTrm) - atom x = do - k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - R.returnTC (RTrm'.atomTrm (k , x)) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - --- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) --- funTypeView = {!!} - -groupoidSolverTerm : Bool → R.Term → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag inp₀ inpF ty = do - - -- inp₀ ← wait-for-type inp₀' - - -- R.typeError (niceAtomList [) - (A' , (t0' , t1')) ← wait-for-type ty >>= (get-boundaryWithDom ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - -- A ← wait-for-type A' - -- t0 ← wait-for-type t0' - -- t1 ← wait-for-type t1' - - let t0 = t0' - let t1 = t1' - - (AA , _) ← get-boundaryWithDom A' >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary2" ]) - (λ x → R.returnTC x) - - let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) - -- (R.typeError [ inp ]ₑ) - - - - [t] ← reflected[Typeₙ]→[reflectedTy] inp - [f] ← reflected[Fns]→[reflectedFn] inpF - - (r0 , r0') ← tryPathE.try≡ inp [t] [f] 100 t0 - (r1 , r1') ← tryPathE.try≡ inp [t] [f] 100 t1 - - - - -- (aL' , (_ , e0)) ← unMapAtoms [] r0' - -- (aL , (_ , e1)) ← unMapAtoms aL' r1' - -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) - show0 ← showRTrmTC r0 - show1 ← showRTrmTC r1 - - red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' - red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' - - - let invPa0 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red0 - invPa1 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red1 - - let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl - ("LHS: ") ∷ₑ show0 ∷nl - ("RHS: ") ∷ₑ show1 ∷nl - (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) - -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - - -- ∷ (R.strErr "\n") - -- (niceAtomList aL) - - let finalTrm0 = - just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) - ∷ invPa0 - - finalTrm1 = - just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) - ∷ invPa1 - - let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms - (finalTrm0 ++ symPathTerms finalTrm1) - - pure (finalTrm , dbgInfo) - - -groupoidSolverMain : Bool → R.Term → R.Term → R.Term → R.TC Unit -groupoidSolverMain debugFlag inp inpF hole = do - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - -- hole' ← R.withNormalisation true $ R.checkType hole ty - (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty - R.catchTC - (R.unify hole solution) - (R.typeError msg) - -groupoidSolverMain' : Bool → R.Term → R.TC Unit -groupoidSolverMain' debugFlag hole = do - - let inpF = R.con (quote [fn]) [] - let inp = R.con (quote [Tω]) [] - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - -- hole' ← R.withNormalisation true $ R.checkType hole ty - - (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty - -- R.typeError [ solution ]ₑ - R.catchTC - (R.unify hole solution) - (R.typeError msg) - - -groupoidSolverMainTerm' : Bool → R.Term → R.Term → R.TC Unit -groupoidSolverMainTerm' debugFlag ty hole = do - - let inpF = R.con (quote [fn]) [] - let inp = R.con (quote [Tω]) [] - - -- hole' ← R.withNormalisation true $ R.checkType hole ty - -- R.commitTC - (solution , msg) ← groupoidSolverTerm debugFlag inp inpF ty - solutionR ← R.quoteTC solution - R.catchTC - (R.unify hole solutionR) - (R.typeError msg) - - - --- -- squareSolverMain : Bool → R.Term → R.Term → R.TC Unit --- -- squareSolverMain debugFlag inp hole = do --- -- ty ← R.inferType hole >>= wait-for-type --- -- hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - --- -- (solution , msg) ← groupoidSolverTerm debugFlag inp hole' --- -- R.catchTC --- -- (R.unify hole' solution) --- -- (R.typeError msg) - --- -- R.catchTC --- -- (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) --- -- (R.typeError [ R.strErr "xxx" ] ) --- -- where --- -- extractMeta : R.Term → R.TC R.Term --- -- extractMeta (R.def _ tl) = matchVarg tl --- -- extractMeta t = --- -- R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) - -macro - solveGroupoidDebug : R.Term → R.Term → R.Term → R.TC Unit - solveGroupoidDebug = groupoidSolverMain true - - ≡! : R.Term → R.TC Unit - ≡! = groupoidSolverMain' true - - ≡!trm : R.Term → R.Term → R.TC Unit - ≡!trm = groupoidSolverMainTerm' true - - --- -- solveSquareDebug : R.Term → R.TC Unit --- -- solveSquareDebug = squareSolverMain false - --- -- solveGroupoid : R.Term → R.TC Unit --- -- solveGroupoid = groupoidSolverMain true - --- -- solveSquare : R.Term → R.TC Unit --- -- solveSquare = squareSolverMain false - - --- module SimpleTest where - --- module _ (n : ℕ) where - --- data AA : Type where --- x y z w : AA --- p p' : x ≡ y --- q : y ≡ z --- q' : z ≡ y --- r : z ≡ w - - - --- pA pB : Path (AA 3) (x) (w) --- pA = ((refl ∙ p) ∙ q) ∙ r --- pB = ((refl ∙ p) ∙ q) ∙ r - - --- pA=pB : Path (Path (AA 3) x y) --- (refl ∙ p) --- (p) --- pA=pB = ≡! - - -module Examples (A : Type ℓ) (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - - pA pB pC : x ≡ w - pA = (p ∙∙ refl ∙∙ q) ∙ sym r - pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - - pA=pB : pA ≡ pB - pA=pB = ≡! - - --- module Examples' (A : ℕ → Type ℓ) --- (x y z w : ∀ {n} → A n) --- (p p' : ∀ {n} → x {n} ≡ y) --- (q : ∀ {n} → y {n} ≡ z) --- (q' : ∀ {n} → z {n} ≡ y) --- (r : ∀ {n} → w {n} ≡ z) where - --- pA pB : Path (A 3) x w --- pA = (p ∙∙ refl ∙∙ q) ∙ sym r --- pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - --- pA=pB : pA ≡ pB --- pA=pB = ≡! - - --- module Examplesℕ' where - --- data AA : (n : ℕ) → Type where --- x y z w : ∀ {n} → AA n --- p p' : ∀ {n} → x {n} ≡ y --- q : ∀ {n} → y {n} ≡ z --- q' : ∀ {n} → z {n} ≡ y --- r : ∀ {n} → w {n} ≡ z - - - --- pA pB : Path (AA 3) (x) (w) --- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) --- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) --- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl - --- pA=pB : pA ≡ pB --- pA=pB = ≡! - - - - --- module Examplesℕ where - --- data AA (n : ℕ) : Type where --- x y z w : AA n --- p p' : x ≡ y --- q : y ≡ z --- q' : z ≡ y --- r : w ≡ z - - - --- pA pB : Path (AA 3) (x) (w) --- pA = (p ∙∙ refl ∙∙ q) ∙ sym (r) --- pB = ((p) ∙ ((q) ∙ sym (sym (r) ∙ (r)) ∙ sym (q)) --- ∙∙ (q) ∙∙ refl) ∙∙ sym (r) ∙∙ refl --- -- -- -- pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - --- pA=pB : R.Term --- pA=pB = {!≡!trm (pA ≡ pB)!} - - - - --- -- -- -- -- -- -- -- -- -- module Examples2 (A B : Type ℓ) (x y z : B) (w : A) (f g : B → A) --- -- -- -- -- -- -- -- -- -- (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ f z) where - --- -- -- -- -- -- -- -- -- -- pA pB : f x ≡ w --- -- -- -- -- -- -- -- -- -- pA = cong f (p ∙∙ refl ∙∙ q) ∙ sym r --- -- -- -- -- -- -- -- -- -- pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl - - --- -- -- -- -- -- -- -- -- -- pA=pB : pA ≡ pB --- -- -- -- -- -- -- -- -- -- pA=pB = solveGroupoidDebug (B ∷Tω A ∷Tω [Tω]) (g ∷fn f ∷fn [fn]) - diff --git a/Cubical/Tactics/Reflection.agda b/Cubical/Tactics/Reflection.agda index f31e4d5134..0e08a7a3f0 100644 --- a/Cubical/Tactics/Reflection.agda +++ b/Cubical/Tactics/Reflection.agda @@ -33,12 +33,52 @@ _>=>_ : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} → (A → TC B) → (B → TC C) → A → TC C (x >=> x₁) x₂ = x x₂ >>= x₁ -wait-for-args : List (Arg Term) → TC (List (Arg Term)) +wait-for-term-args : List (Arg Term) → TC (List (Arg Term)) +wait-for-term-clauses : List (Clause) → TC (List Clause) +wait-for-term-clause : Clause → TC Clause +wait-for-term : Term → TC Term +wait-for-term-telescope : Telescope → TC Telescope + + +wait-for-term (var x args) = var x <$> wait-for-term-args args +wait-for-term (con c args) = con c <$> wait-for-term-args args +wait-for-term (def f args) = def f <$> wait-for-term-args args +wait-for-term (lam v (abs x t)) = (λ t → (lam v (abs x t))) <$> wait-for-term t +wait-for-term (pat-lam cs args) = + ⦇ pat-lam (wait-for-term-clauses cs) (wait-for-term-args args) ⦈ +wait-for-term (pi (arg i a) (abs x b)) = do + a ← wait-for-term a + b ← wait-for-term b + returnTC (pi (arg i a) (abs x b)) +wait-for-term (agda-sort s) = returnTC (agda-sort s) +wait-for-term (lit l) = returnTC (lit l) +wait-for-term (meta x x₁) = blockOnMeta x +wait-for-term unknown = returnTC unknown + +wait-for-term-args [] = ⦇ [] ⦈ +wait-for-term-args (arg i a ∷ xs) = + (_∷_ <$> (arg i <$> wait-for-term a)) <*> wait-for-term-args xs + +wait-for-term-clause (clause tel ps t) = + ⦇ clause (wait-for-term-telescope tel) ⦇ ps ⦈ (wait-for-term t) ⦈ +wait-for-term-clause (absurd-clause tel ps) = + ⦇ absurd-clause (wait-for-term-telescope tel) ⦇ ps ⦈ ⦈ + +wait-for-term-telescope [] = ⦇ [] ⦈ +wait-for-term-telescope ((s , arg i a) ∷ xs) = + ⦇ ⦇ ⦇ s ⦈ , (⦇ arg ⦇ i ⦈ (wait-for-term a) ⦈) ⦈ ∷ (wait-for-term-telescope xs) ⦈ + + +wait-for-term-clauses [] = ⦇ [] ⦈ +wait-for-term-clauses (x ∷ xs) = + ⦇ (wait-for-term-clause x) ∷ wait-for-term-clauses xs ⦈ + +wait-for-type-args : List (Arg Term) → TC (List (Arg Term)) wait-for-type : Term → TC Term -wait-for-type (var x args) = var x <$> wait-for-args args -wait-for-type (con c args) = con c <$> wait-for-args args -wait-for-type (def f args) = def f <$> wait-for-args args +wait-for-type (var x args) = var x <$> wait-for-type-args args +wait-for-type (con c args) = con c <$> wait-for-type-args args +wait-for-type (def f args) = def f <$> wait-for-type-args args wait-for-type (lam v (abs x t)) = returnTC (lam v (abs x t)) wait-for-type (pat-lam cs args) = returnTC (pat-lam cs args) wait-for-type (pi (arg i a) (abs x b)) = do @@ -50,9 +90,9 @@ wait-for-type (lit l) = returnTC (lit l) wait-for-type (meta x x₁) = blockOnMeta x wait-for-type unknown = returnTC unknown -wait-for-args [] = returnTC [] -wait-for-args (arg i a ∷ xs) = - (_∷_ <$> (arg i <$> wait-for-type a)) <*> wait-for-args xs +wait-for-type-args [] = returnTC [] +wait-for-type-args (arg i a ∷ xs) = + (_∷_ <$> (arg i <$> wait-for-type a)) <*> wait-for-type-args xs unapply-path : Term → TC (Maybe (Term × Term × Term)) unapply-path red@(def (quote PathP) (l h∷ T v∷ x v∷ y v∷ [])) = do From ceaac97ec0a2f4941b97117c4509fc753b1b2520 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Tue, 2 Apr 2024 03:55:03 +0200 Subject: [PATCH 12/20] cleanup --- .../EilenbergMacLane/Groups/KleinBottle.agda | 13 +- .../Homotopy/EilenbergMacLane/Properties.agda | 1917 ++++++++--------- 2 files changed, 964 insertions(+), 966 deletions(-) diff --git a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda index 50d46656e9..816f326ec8 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Groups/KleinBottle.agda @@ -47,8 +47,6 @@ open import Cubical.HITs.EilenbergMacLane1 renaming (elimProp to elimPropEM1 ; e open import Cubical.HITs.Susp open import Cubical.HITs.RPn -open import Cubical.Tactics.PathSolver.Solver4 - open AbGroupStr private @@ -361,10 +359,13 @@ fst (H¹[K²,G]≅G×H¹[RP²,G] G) = isoToEquiv mainIso F→ f (square i j) = help i (~ j) where help : cong f (sym line1) ≡ cong f line1 - help = - ≡! - ∙∙ cong (sym (cong f line2) ∙_) (isCommΩEM-base 0 _ _ _) - ∙∙ PathP→compPathL (λ i j → f (square j i)) + help = + lUnit (cong f (sym line1)) + ∙ cong (_∙ cong f (sym line1)) + (sym (lCancel _)) + ∙ sym (assoc _ _ _) + ∙ cong (sym (cong f line2) ∙_) (isCommΩEM-base 0 _ _ _) + ∙ PathP→compPathL (λ i j → f (square j i)) F← : (g : fst G) → (RP² → EM G 1) → KleinBottle → EM G 1 F← g f point = f point diff --git a/Cubical/Homotopy/EilenbergMacLane/Properties.agda b/Cubical/Homotopy/EilenbergMacLane/Properties.agda index bcd227d20a..953ce48309 100644 --- a/Cubical/Homotopy/EilenbergMacLane/Properties.agda +++ b/Cubical/Homotopy/EilenbergMacLane/Properties.agda @@ -38,8 +38,6 @@ open import Cubical.HITs.EilenbergMacLane1 renaming (rec to EMrec ; elim to EM₁elim) open import Cubical.HITs.Susp -open import Cubical.Tactics.PathSolver.Solver4 - private variable ℓ ℓ' ℓ'' : Level @@ -77,7 +75,9 @@ module _ (Ĝ : Group ℓ) where emloop-id : emloop 1g ≡ refl emloop-id = - emloop 1g ≡⟨ zz ⟩ + emloop 1g ≡⟨ rUnit (emloop 1g) ⟩ + emloop 1g ∙ refl ≡⟨ cong (emloop 1g ∙_) (rCancel (emloop 1g) ⁻¹) ⟩ + emloop 1g ∙ (emloop 1g ∙ (emloop 1g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ (emloop 1g ∙ emloop 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (_∙ emloop 1g ⁻¹) ((emloop-comp Ĝ 1g 1g) ⁻¹) ⟩ emloop (1g · 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (λ g → emloop {Group = Ĝ} g @@ -85,10 +85,7 @@ module _ (Ĝ : Group ℓ) where (·IdR 1g) ⟩ emloop 1g ∙ (emloop 1g) ⁻¹ ≡⟨ rCancel (emloop 1g) ⟩ refl ∎ - where - zz : emloop 1g ≡ (emloop 1g ∙ emloop 1g) ∙ emloop 1g ⁻¹ - zz = solveGroupoidDebug ((EM₁ Ĝ) ∷Tω [Tω]) [fn] - + emloop-inv : (g : G) → emloop (inv g) ≡ (emloop g) ⁻¹ emloop-inv g = emloop (inv g) ≡⟨ rUnit (emloop (inv g)) ⟩ @@ -104,956 +101,956 @@ module _ (Ĝ : Group ℓ) where refl ∙ (emloop g) ⁻¹ ≡⟨ (lUnit ((emloop g) ⁻¹)) ⁻¹ ⟩ (emloop g) ⁻¹ ∎ --- isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) --- isGroupoidEM₁ = emsquash - --- --------- Ω (EM₁ G) ≃ G --------- - --- {- since we write composition in diagrammatic order, --- and function composition in the other order, --- we need right multiplication here -} --- rightEquiv : (g : G) → G ≃ G --- rightEquiv g = isoToEquiv isom --- where --- isom : Iso G G --- isom .Iso.fun = _· g --- isom .Iso.inv = _· inv g --- isom .Iso.rightInv h = --- (h · inv g) · g ≡⟨ (·Assoc h (inv g) g) ⁻¹ ⟩ --- h · inv g · g ≡⟨ cong (h ·_) (·InvL g) ⟩ --- h · 1g ≡⟨ ·IdR h ⟩ h ∎ --- isom .Iso.leftInv h = --- (h · g) · inv g ≡⟨ (·Assoc h g (inv g)) ⁻¹ ⟩ --- h · g · inv g ≡⟨ cong (h ·_) (·InvR g) ⟩ --- h · 1g ≡⟨ ·IdR h ⟩ h ∎ - --- compRightEquiv : (g h : G) --- → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) --- compRightEquiv g h = equivEq (funExt (λ x → (·Assoc x g h) ⁻¹)) - --- CodesSet : EM₁ Ĝ → hSet ℓ --- CodesSet = EMrec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp --- where --- RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) --- RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) - --- lemma₁ : (g h : G) → Square --- (ua (rightEquiv g)) (ua (rightEquiv (g · h))) --- refl (ua (rightEquiv h)) --- lemma₁ g h = invEq --- (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) --- refl (ua (rightEquiv h))) --- (ua (rightEquiv g) ∙ ua (rightEquiv h) --- ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ --- ua (compEquiv (rightEquiv g) (rightEquiv h)) --- ≡⟨ cong ua (compRightEquiv g h) ⟩ --- ua (rightEquiv (g · h)) ∎) - --- REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) --- REComp g h = ΣSquareSet (λ _ → isProp→isSet isPropIsSet) (lemma₁ g h) --- Codes : EM₁ Ĝ → Type ℓ --- Codes x = CodesSet x .fst - --- encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x --- encode x p = subst Codes p 1g - --- decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x --- decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) --- emloop λ g → ua→ λ h → emcomp h g - --- decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p --- decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) --- (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ --- emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p - --- encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c --- encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) --- λ g → encode embase (decode embase g) ≡⟨ refl ⟩ --- encode embase (emloop g) ≡⟨ refl ⟩ --- transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ --- 1g · g ≡⟨ ·IdL g ⟩ --- g ∎ - --- ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G --- Iso.fun ΩEM₁Iso = encode embase --- Iso.inv ΩEM₁Iso = emloop --- Iso.rightInv ΩEM₁Iso = encode-decode embase --- Iso.leftInv ΩEM₁Iso = decode-encode embase - --- ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G --- ΩEM₁≡ = isoToPath ΩEM₁Iso - --- --------- Ω (EMₙ₊₁ G) ≃ EMₙ G --------- --- module _ {G : AbGroup ℓ} where --- open AbGroupStr (snd G) --- renaming (_+_ to _+G_ ; -_ to -G_ ; +Assoc to +AssocG) - --- CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓ (3 + n) --- CODE n = --- Trunc.elim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) --- λ { north → EM G (suc n) , hLevelEM G (suc n) --- ; south → EM G (suc n) , hLevelEM G (suc n) --- ; (merid a i) → fib n a i} --- where --- fib : (n : ℕ) → (a : EM-raw G (suc n)) --- → Path (TypeOfHLevel _ (3 + n)) --- (EM G (suc n) , hLevelEM G (suc n)) --- (EM G (suc n) , hLevelEM G (suc n)) --- fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) --- (isoToPath (addIso 1 a)) --- fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) --- (isoToPath (addIso (suc (suc n)) ∣ a ∣)) - --- decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x --- decode' n = --- Trunc.elim (λ _ → isOfHLevelΠ (4 + n) --- λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) --- λ { north → f n --- ; south → g n --- ; (merid a i) → main a i} --- where --- f : (n : ℕ) → _ --- f n = σ-EM' n - --- g : (n : ℕ) → EM G (suc n) → ∣ ptEM-raw ∣ ≡ ∣ south ∣ --- g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptEM-raw) - --- lem₂ : (n : ℕ) (a x : _) --- → Path (Path (EM G (suc (suc n))) _ _) --- ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) --- (g n (EM-raw→EM G (suc n) x)) --- lem₂ zero a x = --- cong (f zero x ∙_) --- (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) --- ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) --- ∙∙ sym (∙assoc _ _ _) --- ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) --- ∙ sym (rUnit _)) --- lem₂ (suc n) a x = --- cong (f (suc n) ∣ x ∣ ∙_) --- ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) --- ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) --- ∙∙ sym (∙assoc _ _ _) --- ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) --- ∙ sym (rUnit _))) - --- lem : (n : ℕ) (x a : EM-raw G (suc n)) --- → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) --- (EM-raw→EM G (suc n) x)) --- ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) --- lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) --- ∙∙ σ-EM'-hom zero x (-ₖ a) --- ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) --- lem (suc n) x a = --- cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) --- ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) --- ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) - --- main : (a : _) --- → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst --- → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) --- main a = --- toPathP (funExt --- (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) --- λ x → --- ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (i ∨ j) ∣) --- i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) --- ∙∙ sym (∙assoc _ _ _) --- ∙∙ lem₂ n a x)) - --- encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst --- encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) - --- decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) --- → decode' n x (encode' n x p) ≡ p --- decode'-encode' zero x = --- J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) --- (σ-EM'-0ₖ 0) --- decode'-encode' (suc n) x = --- J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) --- (σ-EM'-0ₖ (suc n)) - --- encode'-decode' : (n : ℕ) (x : _) --- → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x --- encode'-decode' zero x = --- cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) --- ∙∙ substComposite (λ x → CODE zero x .fst) --- (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) embase --- ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) --- (λ i → lUnitₖ 1 (transportRefl x i) i) --- ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) --- encode'-decode' (suc n) = --- Trunc.elim (λ _ → isOfHLevelTruncPath {n = 4 + n}) --- λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) --- ∙∙ substComposite (λ x → CODE (suc n) x .fst) --- (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) (0ₖ (2 + n)) --- ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptEM-raw (~ i) ∣ₕ)) --- (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) --- ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - --- Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙ G (suc n)))) --- Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) --- Iso.fun (Iso-EM-ΩEM+1 (suc n)) = decode' n (0ₖ (2 + n)) --- Iso.inv (Iso-EM-ΩEM+1 (suc n)) = encode' n ∣ north ∣ --- Iso.rightInv (Iso-EM-ΩEM+1 (suc n)) = decode'-encode' _ _ --- Iso.leftInv (Iso-EM-ΩEM+1 (suc n)) = encode'-decode' _ - --- EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) --- EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) - --- -- Some properties of the isomorphism --- EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) --- EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) - --- ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙ G (suc n))) → EM G n --- ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) - --- EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl --- EM→ΩEM+1-0ₖ zero = emloop-1g _ --- EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) --- EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) - --- EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) --- → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y --- EM→ΩEM+1-hom zero x y = emloop-comp _ x y --- EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y --- EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y - --- EM→ΩEM+1-sym : (n : ℕ) (x : EM G n) → EM→ΩEM+1 n (-ₖ x) ≡ sym (EM→ΩEM+1 n x) --- EM→ΩEM+1-sym n = --- morphLemmas.distrMinus _+ₖ_ _∙_ (EM→ΩEM+1 n) (EM→ΩEM+1-hom n) --- (0ₖ n) refl --- -ₖ_ sym --- (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) --- (lCancelₖ n) rCancel --- ∙assoc --- (EM→ΩEM+1-0ₖ n) - --- ΩEM+1→EM-sym : (n : ℕ) (p : typ (Ω (EM∙ G (suc n)))) --- → ΩEM+1→EM n (sym p) ≡ -ₖ (ΩEM+1→EM n p) --- ΩEM+1→EM-sym n p = sym (cong (ΩEM+1→EM n) (EM→ΩEM+1-sym n (ΩEM+1→EM n p) --- ∙ cong sym (Iso.rightInv (Iso-EM-ΩEM+1 n) p))) --- ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (-ₖ ΩEM+1→EM n p) - --- ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙ G (suc n)))) --- → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) --- ΩEM+1→EM-hom n = --- morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) --- (EM→ΩEM+1-hom n) (ΩEM+1→EM n) --- (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - --- ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n --- ΩEM+1→EM-refl zero = transportRefl 0g --- ΩEM+1→EM-refl (suc zero) = refl --- ΩEM+1→EM-refl (suc (suc n)) = refl - --- EM→ΩEM+1∙ : (n : ℕ) → EM∙ G n →∙ Ω (EM∙ G (suc n)) --- EM→ΩEM+1∙ n .fst = EM→ΩEM+1 n --- EM→ΩEM+1∙ zero .snd = emloop-1g (AbGroup→Group G) --- EM→ΩEM+1∙ (suc zero) .snd = cong (cong ∣_∣ₕ) (rCancel (merid embase)) --- EM→ΩEM+1∙ (suc (suc n)) .snd = cong (cong ∣_∣ₕ) (rCancel (merid north)) - --- EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) --- EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) - --- isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) --- isHomogeneousEM n x = --- ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) - --- isCommΩEM-base : (n : ℕ) (x : _) --- (p q : typ (Ω (EM G (suc n) , x))) → p ∙ q ≡ q ∙ p --- isCommΩEM-base n = --- EM-raw'-elim _ _ (λ _ → isOfHLevelΠ2 (2 + n) --- λ _ _ → isOfHLevelPath (2 + n) --- (hLevelEM _ (suc n) _ _) _ _) --- (EM-raw'-trivElim _ _ --- (λ _ → isOfHLevelΠ2 (suc n) λ _ _ → hLevelEM _ (suc n) _ _ _ _) --- (lem n)) --- where --- * : (n : ℕ) → _ --- * n = EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) - --- lem : (n : ℕ) (p q : typ (Ω (EM G (suc n) , * n))) --- → p ∙ q ≡ q ∙ p --- lem zero = isCommΩEM zero --- lem (suc n) = isCommΩEM (suc n) - --- -- ΩEM+1→EM for arbitrarily based loops. Defining it by pattern --- -- matching is more involved but should give better computational --- -- properties. --- ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) → x ≡ x → EM G n --- ΩEM+1→EM-gen zero = --- elimSet _ (λ _ → isSetΠ λ _ → is-set) (ΩEM+1→EM 0) --- λ g → toPathP (funExt --- λ q → transportRefl (ΩEM+1→EM 0 --- (transport (λ i → Path (EM G (suc zero)) --- (emloop g (~ i)) (emloop g (~ i))) q)) --- ∙ cong (ΩEM+1→EM 0) --- (fromPathP --- (doubleCompPath-filler (emloop g) q (sym (emloop g)) --- ▷ (doubleCompPath≡compPath _ _ _ --- ∙∙ cong (emloop g ∙_) (isCommΩEM 0 q (sym (emloop g))) --- ∙∙ ∙assoc _ _ _ --- ∙∙ cong (_∙ q) (rCancel (emloop g)) --- ∙∙ sym (lUnit q))))) --- ΩEM+1→EM-gen (suc n) = --- Trunc.elim (λ _ → isOfHLevelΠ (4 + n) --- λ _ → isOfHLevelSuc (3 + n) (hLevelEM _ (suc n))) --- λ { north → ΩEM+1→EM (suc n) --- ; south p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) --- ∙∙ p --- ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) --- ; (merid a i) → help a i} --- where --- help : (a : EM-raw G (suc n)) --- → PathP (λ i → Path (EM G (suc (suc n))) ∣ merid a i ∣ ∣ merid a i ∣ --- → EM G (suc n)) --- (ΩEM+1→EM (suc n)) --- λ p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) --- ∙∙ p --- ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) --- help a = --- toPathP (funExt (λ p → --- (transportRefl (ΩEM+1→EM (suc n) --- (transport (λ i → Path (EM G (suc (suc n))) --- ∣ merid a (~ i) ∣ ∣ merid a (~ i) ∣) p)) --- ∙ cong (ΩEM+1→EM (suc n)) --- (flipTransport --- (((rUnit p --- ∙ cong (p ∙_) (sym (lCancel _))) --- ∙ isCommΩEM-base (suc n) ∣ south ∣ p _) --- ∙∙ sym (∙assoc _ _ p) --- ∙∙ cong₂ _∙_ (cong (cong ∣_∣ₕ) --- (sym (cong sym (symDistr --- (sym (merid a)) (merid ptEM-raw))))) --- (isCommΩEM-base _ _ _ p) --- ∙∙ sym (doubleCompPath≡compPath _ _ _) --- ∙∙ λ j → transp (λ i → Path (EM G (suc (suc n))) --- ∣ merid a (i ∨ ~ j) ∣ ∣ merid a (i ∨ ~ j) ∣) (~ j) --- (cong ∣_∣ₕ (compPath-filler' --- (sym (merid a)) (merid ptEM-raw) (~ j)) --- ∙∙ p --- ∙∙ cong ∣_∣ₕ (compPath-filler --- (sym (merid ptEM-raw)) (merid a) (~ j)))))))) - --- EM→ΩEM+1∘EM-raw→EM : (n : ℕ) (x : EM-raw G (suc n)) --- → EM→ΩEM+1 (suc n) (EM-raw→EM _ _ x) ≡ cong ∣_∣ₕ (merid x ∙ sym (merid ptEM-raw)) --- EM→ΩEM+1∘EM-raw→EM zero x = refl --- EM→ΩEM+1∘EM-raw→EM (suc n) x = refl - --- EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) --- → EM G n → x ≡ x --- EM→ΩEM+1-gen n x p = --- sym (rUnitₖ (suc n) x) --- ∙∙ cong (x +ₖ_) (EM→ΩEM+1 n p) --- ∙∙ rUnitₖ (suc n) x - --- ΩEM+1-gen→EM-0ₖ : (n : ℕ) (x : _) --- → ΩEM+1→EM-gen n (0ₖ (suc n)) x --- ≡ ΩEM+1→EM n x --- ΩEM+1-gen→EM-0ₖ zero p = refl --- ΩEM+1-gen→EM-0ₖ (suc n) p = refl - --- EM→ΩEM+1-gen-0ₖ : (n : ℕ) (x : _) --- → EM→ΩEM+1-gen n (0ₖ (suc n)) x --- ≡ EM→ΩEM+1 n x --- EM→ΩEM+1-gen-0ₖ zero x = sym (rUnit _) --- ∙ λ j i → lUnitₖ 1 (EM→ΩEM+1 zero x i) j --- EM→ΩEM+1-gen-0ₖ (suc n) x = sym (rUnit _) --- ∙ λ j i → lUnitₖ (suc (suc n)) (EM→ΩEM+1 (suc n) x i) j - --- EM→ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) --- → (y : EM G n) → ΩEM+1→EM-gen n x (EM→ΩEM+1-gen n x y) ≡ y --- EM→ΩEM+1→EM-gen n = --- EM-raw'-elim _ _ --- (λ _ → isOfHLevelΠ (suc (suc n)) --- (λ _ → isOfHLevelPath (suc (suc n)) --- (hLevelEM _ n) _ _)) --- (EM-raw'-trivElim _ n --- (λ _ → isOfHLevelΠ (suc n) λ _ → hLevelEM _ n _ _) --- λ y → cong (λ p → ΩEM+1→EM-gen n p --- (EM→ΩEM+1-gen n p y)) --- (EM-raw'→EM∙ G (suc n)) --- ∙ (λ i → ΩEM+1-gen→EM-0ₖ n (EM→ΩEM+1-gen-0ₖ n y i) i) --- ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) y) - --- ΩEM+1→EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) --- → (y : x ≡ x) → EM→ΩEM+1-gen n x (ΩEM+1→EM-gen n x y) ≡ y --- ΩEM+1→EM→ΩEM+1-gen n = --- EM-raw'-elim _ _ --- (λ _ → isOfHLevelΠ (suc (suc n)) --- (λ _ → isOfHLevelPath (suc (suc n)) --- (hLevelEM _ (suc n) _ _) _ _)) --- (EM-raw'-trivElim _ n --- (λ _ → isOfHLevelΠ (suc n) --- λ _ → hLevelEM _ (suc n) _ _ _ _) --- (subst (λ p → (y : p ≡ p) --- → EM→ΩEM+1-gen n p (ΩEM+1→EM-gen n p y) ≡ y) --- (sym (EM-raw'→EM∙ _ (suc n))) --- λ p → (λ i → EM→ΩEM+1-gen-0ₖ n (ΩEM+1-gen→EM-0ₖ n p i) i) --- ∙ Iso.rightInv (Iso-EM-ΩEM+1 n) p)) - --- Iso-EM-ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) --- → Iso (EM G n) (x ≡ x) --- Iso.fun (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1-gen n x --- Iso.inv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM-gen n x --- Iso.rightInv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM→ΩEM+1-gen n x --- Iso.leftInv (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1→EM-gen n x - --- ΩEM+1→EM-gen-refl : (n : ℕ) (x : EM G (suc n)) --- → ΩEM+1→EM-gen n x refl ≡ 0ₖ n --- ΩEM+1→EM-gen-refl n = --- EM-raw'-elim _ (suc n) --- (λ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _) --- (EM-raw'-trivElim _ n --- (λ _ → hLevelEM _ n _ _) --- (lem n)) --- where --- lem : (n : ℕ) → ΩEM+1→EM-gen n --- (EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) refl --- ≡ 0ₖ n --- lem zero = ΩEM+1→EM-refl 0 --- lem (suc n) = ΩEM+1→EM-refl (suc n) - --- ΩEM+1→EM-gen-hom : (n : ℕ) (x : EM G (suc n)) (p q : x ≡ x) --- → ΩEM+1→EM-gen n x (p ∙ q) ≡ ΩEM+1→EM-gen n x p +ₖ ΩEM+1→EM-gen n x q --- ΩEM+1→EM-gen-hom n = --- EM-raw'-elim _ (suc n) --- (λ _ → isOfHLevelΠ2 (suc (suc n)) --- (λ _ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _)) --- (EM-raw'-trivElim _ n --- (λ _ → isOfHLevelΠ2 (suc n) (λ _ _ → hLevelEM _ n _ _)) --- (lem n)) --- where --- lem : (n : ℕ) (p q : EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) --- ≡ EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) --- → ΩEM+1→EM-gen n _ (p ∙ q) --- ≡ ΩEM+1→EM-gen n _ p --- +ₖ ΩEM+1→EM-gen n _ q --- lem zero p q = ΩEM+1→EM-hom 0 p q --- lem (suc n) p q = ΩEM+1→EM-hom (suc n) p q - --- -- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for --- -- the cup product --- module _ where --- open AbGroupStr renaming (_+_ to comp) - --- isContr-↓∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ H n) --- fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl --- snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = --- Σ≡Prop (λ x → is-set (snd H) _ _) --- (funExt (raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) --- fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl --- fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = --- Trunc.elim {B = λ x → 0ₖ (suc n) ≡ fst f x} --- (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) --- (raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) --- x i --- snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - --- isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) --- → isContr ((EM-raw G (suc n) , ptEM-raw) →∙ EM∙ H n) --- isContr-↓∙' zero = isContr-↓∙ zero --- fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl --- fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = --- raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} --- (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i --- snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) - --- isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) --- → isOfHLevel (suc m) (A →∙ EM∙ G n) --- → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) --- isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ --- where --- step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) --- step₁ = --- subst (isOfHLevel (suc m)) --- (λ i → A →∙ ua∙ {A = Ω (EM∙ G (suc n))} {B = EM∙ G n} --- (invEquiv (EM≃ΩEM+1 n)) --- (ΩEM+1→EM-refl n) (~ i)) hlev - --- step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ EM∙ G (suc n) ∙))) --- step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ - --- step₃ : isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) --- step₃ = isOfHLevelΩ→isOfHLevel m --- λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) --- (isHomogeneous→∙ (isHomogeneousEM (suc n)) f) --- step₂ - --- isOfHLevel↑∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- → ∀ n m → isOfHLevel (2 + n) (EM∙ G (suc m) →∙ EM∙ H (suc (n + m))) --- isOfHLevel↑∙ zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙ m)) --- isOfHLevel↑∙ (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) - --- isOfHLevel↑∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- → ∀ n m → isOfHLevel (2 + n) (EM-raw∙ G (suc m) →∙ EM∙ H (suc (n + m))) --- isOfHLevel↑∙-lem zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) --- isOfHLevel↑∙-lem (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙-lem n m) - --- EM₁→∙Iso : {G : AbGroup ℓ} {H : AbGroup ℓ'} (m : ℕ) --- → Iso (EM-raw'∙ G 1 →∙ EM∙ H (suc m)) (fst G → typ (Ω (EM∙ H (suc m)))) --- Iso.fun (EM₁→∙Iso m) f g = sym (snd f) ∙∙ cong (fst f) (emloop-raw g) ∙∙ snd f --- fst (Iso.inv (EM₁→∙Iso m) f) embase-raw = 0ₖ (suc m) --- fst (Iso.inv (EM₁→∙Iso m) f) (emloop-raw g i) = f g i --- snd (Iso.inv (EM₁→∙Iso m) f) = refl --- Iso.rightInv (EM₁→∙Iso m) f = funExt λ x → sym (rUnit _) --- Iso.leftInv (EM₁→∙Iso m) (f , p) = --- →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt λ { embase-raw → sym p --- ; (emloop-raw g i) j --- → doubleCompPath-filler (sym p) (cong f (emloop-raw g)) p (~ j) i}) - --- isOfHLevel↑∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- → ∀ n m → isOfHLevel (2 + n) (EM-raw'∙ G m →∙ EM∙ H (n + m)) --- isOfHLevel↑∙' {H = H} zero zero = --- isOfHLevelΣ 2 (isOfHLevelΠ 2 (λ _ → AbGroupStr.is-set (snd H))) --- λ _ → isOfHLevelPath 2 (AbGroupStr.is-set (snd H)) _ _ --- isOfHLevel↑∙' zero (suc zero) = --- subst (isOfHLevel 2) (sym (isoToPath (EM₁→∙Iso 0))) --- (isSetΠ λ _ → emsquash _ _) --- isOfHLevel↑∙' zero (suc (suc m)) = isOfHLevel↑∙-lem zero (suc m) --- isOfHLevel↑∙' {H = H} (suc n) zero = --- isOfHLevelΣ (2 + suc n) (isOfHLevelΠ (2 + suc n) --- (λ _ → subst (isOfHLevel (suc (suc (suc n)))) --- (cong (EM H) (cong suc (+-comm 0 n))) --- (hLevelEM H (suc n)))) --- λ _ → isOfHLevelPath (suc (suc (suc n))) --- (subst (isOfHLevel (suc (suc (suc n)))) --- (cong (EM H) (cong suc (+-comm 0 n))) --- (hLevelEM H (suc n))) _ _ --- isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc zero) = --- subst (isOfHLevel (2 + suc n)) (sym (isoToPath (EM₁→∙Iso (suc n))) --- ∙ λ i → EM-raw'∙ G 1 →∙ EM∙ H (suc (+-comm 1 n i))) --- (isOfHLevelΠ (2 + suc n) λ x → (isOfHLevelTrunc (4 + n) _ _)) --- isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc (suc m)) = --- subst (isOfHLevel (2 + suc n)) --- (λ i → (EM-raw'∙ G (suc (suc m)) --- →∙ EM∙ H (suc (+-suc n (suc m) (~ i))))) --- (isOfHLevel↑∙-lem (suc n) (suc m)) - --- →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) --- → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) --- →∙EMPath {G = G} A n = --- ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl --- ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) - --- private --- contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) --- → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) --- fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl --- snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = --- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) --- (sym (funExt (help n f p))) --- where --- help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) --- → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) --- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) --- help zero f p = --- raw-elim G zero --- (λ _ → isOfHLevel↑∙ zero m _ _) p --- help (suc n) f p = --- Trunc.elim (λ _ → isOfHLevelPath (4 + n) --- (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) --- (λ i → suc (suc (+-comm (suc n) 1 i))) --- (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) --- (raw-elim G (suc n) --- (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) - --- contr∙-lem' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) --- → isContr (EM∙ G (suc n) →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) --- fst (contr∙-lem' n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl --- snd (contr∙-lem' {G = G} {H = H} {L = L} n m) (f , p) = --- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) --- (funExt λ x → sym (help' n f p x)) --- where --- help' : (n : ℕ) → (f : EM G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) --- → f (snd (EM∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) --- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) --- help' zero f p = --- raw-elim _ zero (λ _ → isOfHLevel↑∙' zero (suc m) _ _) p --- help' (suc n) f p = --- Trunc.elim (λ _ → isOfHLevelPath (4 + n) --- (subst2 (λ x y → isOfHLevel x (EM-raw'∙ H (suc m) →∙ EM∙ L y)) --- (λ i → suc (suc (suc (+-comm n 1 i)))) --- (cong suc (+-suc n m)) --- (isOfHLevelPlus' {n = 1} (suc (suc (suc n))) --- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m)))) _ _) --- (raw-elim _ (suc n) (λ _ → isOfHLevelPath' (2 + n) --- (subst (λ y → isOfHLevel (suc (suc (suc n))) --- (EM-raw'∙ H (suc m) →∙ EM∙ L y)) --- (+-suc (suc n) m) --- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p) - --- contr∙-lem'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) --- → isContr (EM-raw'∙ G (suc n) --- →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) --- fst (contr∙-lem'' n m) = (λ _ → (λ _ → 0ₖ (suc (n + m))) , refl) , refl --- snd (contr∙-lem'' {G = G} {H = H} {L = L} n m) (f , p) = --- →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) --- (funExt λ x → sym (help n f p x)) --- where --- help : (n : ℕ) → (f : EM-raw' G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) --- → f (snd (EM-raw'∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) --- → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) --- help zero f p = --- EM-raw'-trivElim G zero (λ _ → isOfHLevel↑∙' _ _ _ _) p --- help (suc n) f p = --- EM-raw'-trivElim _ _ (λ _ → isOfHLevelPath' (suc (suc n)) --- (subst (λ y → isOfHLevel (suc (suc (suc n))) (EM-raw'∙ H (suc m) →∙ EM∙ L y)) --- (cong suc (+-suc n m)) --- (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p - --- isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} --- → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) --- →∙ (EM∙ H (suc m) --- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) --- isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m zero = --- isOfHLevelΩ→isOfHLevel 0 λ f --- → subst --- isProp (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) --- where --- h : isProp (EM∙ G (suc n) --- →∙ (Ω (EM∙ H (suc m) --- →∙ EM∙ L (suc (suc (n + m))) ∙))) --- h = --- subst isProp --- (λ i → EM∙ G (suc n) --- →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) --- (isContr→isProp (contr∙-lem n m)) --- isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = --- isOfHLevelΩ→isOfHLevel (suc l) --- λ f → --- subst (isOfHLevel (2 + l)) --- (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) --- where --- h : isOfHLevel (2 + l) --- (EM∙ G (suc n) --- →∙ (Ω (EM∙ H (suc m) --- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) --- h = --- subst (isOfHLevel (2 + l)) --- (λ i → EM∙ G (suc n) --- →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) --- (isOfHLevel↑∙∙ n m l) - --- isOfHLevel↑∙∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} --- → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) --- →∙ (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) --- isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m zero = --- isOfHLevelΩ→isOfHLevel 0 --- λ f → subst --- isProp (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) --- lem) --- where --- lem : isProp (EM∙ G (suc n) --- →∙ (Ω (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (n + m))) ∙))) --- lem = subst isProp --- (λ i → EM∙ G (suc n) --- →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) --- (isContr→isProp (contr∙-lem' n m)) --- isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m (suc l) = --- isOfHLevelΩ→isOfHLevel (suc l) --- λ f → --- subst (isOfHLevel (2 + l)) --- (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) --- where --- lem : isOfHLevel (2 + l) --- (EM∙ G (suc n) --- →∙ (Ω (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) --- lem = --- subst (isOfHLevel (2 + l)) --- (λ i → EM∙ G (suc n) --- →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) --- (isOfHLevel↑∙∙' n m l) - --- isOfHLevel↑∙∙'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} --- → ∀ n m l → isOfHLevel (2 + l) (EM-raw'∙ G (suc n) --- →∙ (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (l + n + m))) ∙)) --- isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m zero = --- isOfHLevelΩ→isOfHLevel 0 --- λ f → subst --- isProp (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) --- lem) --- where --- lem : isProp (EM-raw'∙ G (suc n) --- →∙ (Ω (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (n + m))) ∙))) --- lem = subst isProp --- (λ i → EM-raw'∙ G (suc n) --- →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) --- (isContr→isProp (contr∙-lem'' _ _)) --- isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m (suc l) = --- isOfHLevelΩ→isOfHLevel (suc l) --- λ f → --- subst (isOfHLevel (2 + l)) --- (cong (λ x → typ (Ω x)) --- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) --- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) --- where --- lem : isOfHLevel (2 + l) --- (EM-raw'∙ G (suc n) --- →∙ (Ω (EM-raw'∙ H (suc m) --- →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) --- lem = --- subst (isOfHLevel (2 + l)) --- (λ i → EM-raw'∙ G (suc n) --- →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) --- (isOfHLevel↑∙∙'' n m l) - --- -- A homomorphism φ : G → H of AbGroups induces a homomorphism --- -- φ' : K(G,n) → K(H,n) --- inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- → AbGroupHom G' H' --- → ∀ n --- → EM-raw G' n → EM-raw H' n --- inducedFun-EM-raw f = --- elim+2 (fst f) --- (EMrec _ emsquash embase --- (λ g → emloop (fst f g)) --- λ g h → compPathR→PathP (sym --- (sym (lUnit _) --- ∙∙ cong (_∙ (sym (emloop (fst f h)))) --- (cong emloop (IsGroupHom.pres· (snd f) g h) --- ∙ emloop-comp _ (fst f g) (fst f h)) --- ∙∙ sym (∙assoc _ _ _) --- ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) --- ∙∙ sym (rUnit _)))) --- (λ n ind → λ { north → north --- ; south → south --- ; (merid a i) → merid (ind a) i} ) - --- inducedFun-EM-raw-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM-raw G' n) --- → inducedFun-EM-raw (idGroupHom {G = AbGroup→Group G'}) n x ≡ x --- inducedFun-EM-raw-id zero x = refl --- inducedFun-EM-raw-id (suc zero) = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) --- λ { embase-raw → refl ; (emloop-raw g i) → refl} --- inducedFun-EM-raw-id (suc (suc n)) north = refl --- inducedFun-EM-raw-id (suc (suc n)) south = refl --- inducedFun-EM-raw-id (suc (suc n)) (merid a i) j = merid (inducedFun-EM-raw-id (suc n) a j) i - --- inducedFun-EM-raw-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} --- (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) --- → (x : EM-raw G' n) → inducedFun-EM-raw (compGroupHom ϕ ψ) n x --- ≡ inducedFun-EM-raw ψ n (inducedFun-EM-raw ϕ n x) --- inducedFun-EM-raw-comp ϕ ψ zero x = refl --- inducedFun-EM-raw-comp ϕ ψ (suc zero) = --- EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) --- λ { embase-raw → refl ; (emloop-raw g i) → refl} --- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) north = refl --- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) south = refl --- inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) (merid a i) j = --- merid (inducedFun-EM-raw-comp ϕ ψ (suc n) a j) i - --- inducedFun-EM : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- → AbGroupHom G' H' --- → ∀ n --- → EM G' n → EM H' n --- inducedFun-EM f zero = inducedFun-EM-raw f zero --- inducedFun-EM f (suc zero) = inducedFun-EM-raw f (suc zero) --- inducedFun-EM f (suc (suc n)) = Trunc.map (inducedFun-EM-raw f (2 + n)) - --- EM-raw→EM-funct : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- (n : ℕ) (ψ : AbGroupHom G H) (y : EM-raw G n) --- → EM-raw→EM _ _ (inducedFun-EM-raw ψ n y) --- ≡ inducedFun-EM ψ n (EM-raw→EM _ _ y) --- EM-raw→EM-funct zero ψ y = refl --- EM-raw→EM-funct (suc zero) ψ y = refl --- EM-raw→EM-funct (suc (suc n)) ψ y = refl - --- inducedFun-EM-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM G' n) --- → inducedFun-EM (idGroupHom {G = AbGroup→Group G'}) n x ≡ x --- inducedFun-EM-id zero x = refl --- inducedFun-EM-id (suc zero) = inducedFun-EM-raw-id (suc zero) --- inducedFun-EM-id (suc (suc n)) = --- Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) --- λ x → cong ∣_∣ₕ (inducedFun-EM-raw-id _ x) - --- inducedFun-EM-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} --- (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) --- → (x : EM G' n) → inducedFun-EM (compGroupHom ϕ ψ) n x --- ≡ inducedFun-EM ψ n (inducedFun-EM ϕ n x) --- inducedFun-EM-comp ϕ ψ zero x = refl --- inducedFun-EM-comp ϕ ψ (suc zero) = inducedFun-EM-raw-comp ϕ ψ (suc zero) --- inducedFun-EM-comp ϕ ψ (suc (suc n)) = --- Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) --- λ x → cong ∣_∣ₕ (inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) x) - --- inducedFun-EM0ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {ϕ : AbGroupHom G' H'} (n : ℕ) --- → inducedFun-EM ϕ n (0ₖ n) ≡ 0ₖ n --- inducedFun-EM0ₖ {ϕ = ϕ} zero = IsGroupHom.pres1 (snd ϕ) --- inducedFun-EM0ₖ (suc zero) = refl --- inducedFun-EM0ₖ (suc (suc n)) = refl - --- inducedFun-EM-pres+ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- (ϕ : AbGroupHom G' H') (n : ℕ) (x y : EM G' n) --- → inducedFun-EM ϕ n (x +ₖ y) ≡ inducedFun-EM ϕ n x +ₖ inducedFun-EM ϕ n y --- inducedFun-EM-pres+ₖ ϕ zero x y = IsGroupHom.pres· (snd ϕ) x y --- inducedFun-EM-pres+ₖ {G' = G'} {H' = H'} ϕ (suc n) = --- EM-elim2 (suc n) (λ _ _ → isOfHLevelPath (2 + suc n) (hLevelEM _ (suc n)) _ _) --- (wedgeConEM.fun _ _ n n --- (λ _ _ → isOfHLevelPath' (suc n + suc n) --- (subst (λ m → isOfHLevel (suc (suc m)) (EM H' (suc n))) --- (sym (+-suc n n)) --- (isOfHLevelPlus' {n = n} (3 + n) --- (hLevelEM _ (suc n)))) --- _ _) --- (l n) --- (r n) --- (l≡r n)) --- where --- lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) → EM-raw→EM G (suc n) ptEM-raw ≡ 0ₖ _ --- lem zero = refl --- lem (suc n) = refl - --- l : (n : ℕ) (y : EM-raw G' (suc n)) --- → inducedFun-EM ϕ (suc n) ((EM-raw→EM G' (suc n) ptEM-raw) --- +ₖ EM-raw→EM G' (suc n) y) --- ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) --- +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y) --- l n y = (cong (inducedFun-EM ϕ (suc n)) --- (cong (_+ₖ EM-raw→EM G' (suc n) y) (lem n) --- ∙ lUnitₖ (suc n) (EM-raw→EM G' (suc n) y))) --- ∙∙ sym (lUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) --- ∙∙ cong (_+ₖ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) --- (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) --- ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) - --- r : (n : ℕ) (x : EM-raw G' (suc n)) --- → inducedFun-EM ϕ (suc n) --- (EM-raw→EM G' (suc n) x +ₖ EM-raw→EM G' (suc n) ptEM-raw) --- ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) --- +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) --- r n x = cong (inducedFun-EM ϕ (suc n)) --- (cong (EM-raw→EM G' (suc n) x +ₖ_) (lem n) --- ∙ rUnitₖ (suc n) (EM-raw→EM G' (suc n) x)) --- ∙∙ sym (rUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x))) --- ∙∙ cong (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) +ₖ_) --- (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) --- ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) - --- l≡r : (n : ℕ) → l n ptEM-raw ≡ r n ptEM-raw --- l≡r zero = refl --- l≡r (suc n) = refl - --- inducedFun-EM-pres-ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- (ϕ : AbGroupHom G' H') (n : ℕ) (x : EM G' n) --- → inducedFun-EM ϕ n (-ₖ x) ≡ -ₖ (inducedFun-EM ϕ n x) --- inducedFun-EM-pres-ₖ ϕ n = --- morphLemmas.distrMinus _+ₖ_ _+ₖ_ --- (inducedFun-EM ϕ n) (inducedFun-EM-pres+ₖ ϕ n) (0ₖ n) (0ₖ n) (-ₖ_) (-ₖ_) --- (lUnitₖ n) (rUnitₖ n) --- (lCancelₖ n) (rCancelₖ n) --- (assocₖ n) (inducedFun-EM0ₖ n) - --- EMFun-EM→ΩEM+1 : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- {ϕ : AbGroupHom G H} (n : ℕ) (x : EM G n) --- → PathP (λ i → inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i) --- ≡ inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i)) --- (EM→ΩEM+1 n (inducedFun-EM ϕ n x)) --- (cong (inducedFun-EM ϕ (suc n)) (EM→ΩEM+1 n x)) --- EMFun-EM→ΩEM+1 {ϕ = ϕ} zero x = refl --- EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc zero) x = --- cong-∙ ∣_∣ₕ (merid (inducedFun-EM ϕ (suc zero) x)) --- (sym (merid embase)) --- ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc zero))) --- (cong ∣_∣ₕ (merid x)) (cong ∣_∣ₕ (sym (merid embase)))) --- ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc zero)))) --- (sym (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase)))) --- EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc (suc n)) = --- Trunc.elim (λ _ → isOfHLevelPath (4 + n) --- (isOfHLevelTrunc (5 + n) _ _) _ _) --- λ a → cong-∙ ∣_∣ₕ (merid (inducedFun-EM-raw ϕ (2 + n) a)) --- (sym (merid north)) --- ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc (suc n)))) --- (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (sym (merid north)))) --- ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc (suc n))))) --- (sym (cong-∙ ∣_∣ₕ (merid a) (sym (merid north)))) - - --- inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- → AbGroupEquiv G' H' --- → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) --- Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n --- Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n --- Iso.rightInv (inducedFun-EM-rawIso e n) = h n --- where --- h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) --- (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) --- h = elim+2 --- (secEq (fst e)) --- (elimSet _ (λ _ → emsquash _ _) refl --- (λ g → compPathR→PathP --- (sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) --- (sym (lUnit _)) --- ∙ rCancel _)))) --- λ n p → λ { north → refl --- ; south → refl --- ; (merid a i) k → merid (p a k) i} --- Iso.leftInv (inducedFun-EM-rawIso e n) = h n --- where --- h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) --- (Iso.inv (inducedFun-EM-rawIso e n)) --- h = elim+2 --- (retEq (fst e)) --- (elimSet _ (λ _ → emsquash _ _) refl --- (λ g → compPathR→PathP --- ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) --- ∙ rCancel _))))) --- λ n p → λ { north → refl --- ; south → refl --- ; (merid a i) k → merid (p a k) i} - --- module _ {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) where --- Iso→EMIso : ∀ n → Iso (EM G n) (EM H n) --- Iso.fun (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom e) n --- Iso.inv (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom (invGroupEquiv e)) n --- Iso.rightInv (Iso→EMIso zero) = Iso.rightInv (inducedFun-EM-rawIso e zero) --- Iso.rightInv (Iso→EMIso (suc zero)) = --- Iso.rightInv (inducedFun-EM-rawIso e (suc zero)) --- Iso.rightInv (Iso→EMIso (suc (suc n))) = --- Iso.rightInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) --- Iso.leftInv (Iso→EMIso zero) = --- Iso.leftInv (inducedFun-EM-rawIso e zero) --- Iso.leftInv (Iso→EMIso (suc zero)) = --- Iso.leftInv (inducedFun-EM-rawIso e (suc zero)) --- Iso.leftInv (Iso→EMIso (suc (suc n))) = --- Iso.leftInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) - --- Iso→EMIso∙ : ∀ n → Iso.fun (Iso→EMIso n) (EM∙ G n .snd) ≡ EM∙ H n .snd --- Iso→EMIso∙ zero = IsGroupHom.pres1 (e .snd) --- Iso→EMIso∙ (suc zero) = refl --- Iso→EMIso∙ (suc (suc n)) = refl - --- Iso→EMIso⁻∙ : ∀ n → Iso.inv (Iso→EMIso n) (EM∙ H n .snd) ≡ EM∙ G n .snd --- Iso→EMIso⁻∙ zero = IsGroupHom.pres1 (invGroupEquiv e .snd) --- Iso→EMIso⁻∙ (suc zero) = refl --- Iso→EMIso⁻∙ (suc (suc n)) = refl - --- Iso→EMIsoInv : {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) --- → ∀ n → Iso.inv (Iso→EMIso e n) ≡ Iso.fun (Iso→EMIso (invGroupEquiv e) n) --- Iso→EMIsoInv e zero = refl --- Iso→EMIsoInv e (suc zero) = refl --- Iso→EMIsoInv e (suc (suc n)) = refl - --- EM⊗-commIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} --- → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) --- EM⊗-commIso {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) - --- EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} --- → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) --- EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) + isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) + isGroupoidEM₁ = emsquash + + --------- Ω (EM₁ G) ≃ G --------- + + {- since we write composition in diagrammatic order, + and function composition in the other order, + we need right multiplication here -} + rightEquiv : (g : G) → G ≃ G + rightEquiv g = isoToEquiv isom + where + isom : Iso G G + isom .Iso.fun = _· g + isom .Iso.inv = _· inv g + isom .Iso.rightInv h = + (h · inv g) · g ≡⟨ (·Assoc h (inv g) g) ⁻¹ ⟩ + h · inv g · g ≡⟨ cong (h ·_) (·InvL g) ⟩ + h · 1g ≡⟨ ·IdR h ⟩ h ∎ + isom .Iso.leftInv h = + (h · g) · inv g ≡⟨ (·Assoc h g (inv g)) ⁻¹ ⟩ + h · g · inv g ≡⟨ cong (h ·_) (·InvR g) ⟩ + h · 1g ≡⟨ ·IdR h ⟩ h ∎ + + compRightEquiv : (g h : G) + → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) + compRightEquiv g h = equivEq (funExt (λ x → (·Assoc x g h) ⁻¹)) + + CodesSet : EM₁ Ĝ → hSet ℓ + CodesSet = EMrec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp + where + RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) + RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) + + lemma₁ : (g h : G) → Square + (ua (rightEquiv g)) (ua (rightEquiv (g · h))) + refl (ua (rightEquiv h)) + lemma₁ g h = invEq + (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) + refl (ua (rightEquiv h))) + (ua (rightEquiv g) ∙ ua (rightEquiv h) + ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ + ua (compEquiv (rightEquiv g) (rightEquiv h)) + ≡⟨ cong ua (compRightEquiv g h) ⟩ + ua (rightEquiv (g · h)) ∎) + + REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) + REComp g h = ΣSquareSet (λ _ → isProp→isSet isPropIsSet) (lemma₁ g h) + Codes : EM₁ Ĝ → Type ℓ + Codes x = CodesSet x .fst + + encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x + encode x p = subst Codes p 1g + + decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x + decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) + emloop λ g → ua→ λ h → emcomp h g + + decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p + decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) + (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ + emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p + + encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c + encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) + λ g → encode embase (decode embase g) ≡⟨ refl ⟩ + encode embase (emloop g) ≡⟨ refl ⟩ + transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ + 1g · g ≡⟨ ·IdL g ⟩ + g ∎ + + ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G + Iso.fun ΩEM₁Iso = encode embase + Iso.inv ΩEM₁Iso = emloop + Iso.rightInv ΩEM₁Iso = encode-decode embase + Iso.leftInv ΩEM₁Iso = decode-encode embase + + ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G + ΩEM₁≡ = isoToPath ΩEM₁Iso + +--------- Ω (EMₙ₊₁ G) ≃ EMₙ G --------- +module _ {G : AbGroup ℓ} where + open AbGroupStr (snd G) + renaming (_+_ to _+G_ ; -_ to -G_ ; +Assoc to +AssocG) + + CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓ (3 + n) + CODE n = + Trunc.elim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) + λ { north → EM G (suc n) , hLevelEM G (suc n) + ; south → EM G (suc n) , hLevelEM G (suc n) + ; (merid a i) → fib n a i} + where + fib : (n : ℕ) → (a : EM-raw G (suc n)) + → Path (TypeOfHLevel _ (3 + n)) + (EM G (suc n) , hLevelEM G (suc n)) + (EM G (suc n) , hLevelEM G (suc n)) + fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) + (isoToPath (addIso 1 a)) + fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) + (isoToPath (addIso (suc (suc n)) ∣ a ∣)) + + decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x + decode' n = + Trunc.elim (λ _ → isOfHLevelΠ (4 + n) + λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → f n + ; south → g n + ; (merid a i) → main a i} + where + f : (n : ℕ) → _ + f n = σ-EM' n + + g : (n : ℕ) → EM G (suc n) → ∣ ptEM-raw ∣ ≡ ∣ south ∣ + g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptEM-raw) + + lem₂ : (n : ℕ) (a x : _) + → Path (Path (EM G (suc (suc n))) _ _) + ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) + (g n (EM-raw→EM G (suc n) x)) + lem₂ zero a x = + cong (f zero x ∙_) + (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) + ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _)) + lem₂ (suc n) a x = + cong (f (suc n) ∣ x ∣ ∙_) + ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) + ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _))) + + lem : (n : ℕ) (x a : EM-raw G (suc n)) + → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) + (EM-raw→EM G (suc n) x)) + ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) + lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) + ∙∙ σ-EM'-hom zero x (-ₖ a) + ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) + lem (suc n) x a = + cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) + ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) + ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) + + main : (a : _) + → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst + → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) + main a = + toPathP (funExt + (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) + λ x → + ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (i ∨ j) ∣) + i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) + ∙∙ sym (∙assoc _ _ _) + ∙∙ lem₂ n a x)) + + encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst + encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) + + decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) + → decode' n x (encode' n x p) ≡ p + decode'-encode' zero x = + J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) + (σ-EM'-0ₖ 0) + decode'-encode' (suc n) x = + J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) + (σ-EM'-0ₖ (suc n)) + + encode'-decode' : (n : ℕ) (x : _) + → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x + encode'-decode' zero x = + cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) + ∙∙ substComposite (λ x → CODE zero x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) embase + ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) + (λ i → lUnitₖ 1 (transportRefl x i) i) + ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) + encode'-decode' (suc n) = + Trunc.elim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) + ∙∙ substComposite (λ x → CODE (suc n) x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) (0ₖ (2 + n)) + ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptEM-raw (~ i) ∣ₕ)) + (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + + Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙ G (suc n)))) + Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) + Iso.fun (Iso-EM-ΩEM+1 (suc n)) = decode' n (0ₖ (2 + n)) + Iso.inv (Iso-EM-ΩEM+1 (suc n)) = encode' n ∣ north ∣ + Iso.rightInv (Iso-EM-ΩEM+1 (suc n)) = decode'-encode' _ _ + Iso.leftInv (Iso-EM-ΩEM+1 (suc n)) = encode'-decode' _ + + EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) + EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) + + -- Some properties of the isomorphism + EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) + EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) + + ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙ G (suc n))) → EM G n + ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + + EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl + EM→ΩEM+1-0ₖ zero = emloop-1g _ + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) + + EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) + → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y + EM→ΩEM+1-hom zero x y = emloop-comp _ x y + EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y + EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y + + EM→ΩEM+1-sym : (n : ℕ) (x : EM G n) → EM→ΩEM+1 n (-ₖ x) ≡ sym (EM→ΩEM+1 n x) + EM→ΩEM+1-sym n = + morphLemmas.distrMinus _+ₖ_ _∙_ (EM→ΩEM+1 n) (EM→ΩEM+1-hom n) + (0ₖ n) refl + -ₖ_ sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (lCancelₖ n) rCancel + ∙assoc + (EM→ΩEM+1-0ₖ n) + + ΩEM+1→EM-sym : (n : ℕ) (p : typ (Ω (EM∙ G (suc n)))) + → ΩEM+1→EM n (sym p) ≡ -ₖ (ΩEM+1→EM n p) + ΩEM+1→EM-sym n p = sym (cong (ΩEM+1→EM n) (EM→ΩEM+1-sym n (ΩEM+1→EM n p) + ∙ cong sym (Iso.rightInv (Iso-EM-ΩEM+1 n) p))) + ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) (-ₖ ΩEM+1→EM n p) + + ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙ G (suc n)))) + → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) + ΩEM+1→EM-hom n = + morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) + (EM→ΩEM+1-hom n) (ΩEM+1→EM n) + (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) + + ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n + ΩEM+1→EM-refl zero = transportRefl 0g + ΩEM+1→EM-refl (suc zero) = refl + ΩEM+1→EM-refl (suc (suc n)) = refl + + EM→ΩEM+1∙ : (n : ℕ) → EM∙ G n →∙ Ω (EM∙ G (suc n)) + EM→ΩEM+1∙ n .fst = EM→ΩEM+1 n + EM→ΩEM+1∙ zero .snd = emloop-1g (AbGroup→Group G) + EM→ΩEM+1∙ (suc zero) .snd = cong (cong ∣_∣ₕ) (rCancel (merid embase)) + EM→ΩEM+1∙ (suc (suc n)) .snd = cong (cong ∣_∣ₕ) (rCancel (merid north)) + + EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) + EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) + + isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) + isHomogeneousEM n x = + ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) + + isCommΩEM-base : (n : ℕ) (x : _) + (p q : typ (Ω (EM G (suc n) , x))) → p ∙ q ≡ q ∙ p + isCommΩEM-base n = + EM-raw'-elim _ _ (λ _ → isOfHLevelΠ2 (2 + n) + λ _ _ → isOfHLevelPath (2 + n) + (hLevelEM _ (suc n) _ _) _ _) + (EM-raw'-trivElim _ _ + (λ _ → isOfHLevelΠ2 (suc n) λ _ _ → hLevelEM _ (suc n) _ _ _ _) + (lem n)) + where + * : (n : ℕ) → _ + * n = EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) + + lem : (n : ℕ) (p q : typ (Ω (EM G (suc n) , * n))) + → p ∙ q ≡ q ∙ p + lem zero = isCommΩEM zero + lem (suc n) = isCommΩEM (suc n) + + -- ΩEM+1→EM for arbitrarily based loops. Defining it by pattern + -- matching is more involved but should give better computational + -- properties. + ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) → x ≡ x → EM G n + ΩEM+1→EM-gen zero = + elimSet _ (λ _ → isSetΠ λ _ → is-set) (ΩEM+1→EM 0) + λ g → toPathP (funExt + λ q → transportRefl (ΩEM+1→EM 0 + (transport (λ i → Path (EM G (suc zero)) + (emloop g (~ i)) (emloop g (~ i))) q)) + ∙ cong (ΩEM+1→EM 0) + (fromPathP + (doubleCompPath-filler (emloop g) q (sym (emloop g)) + ▷ (doubleCompPath≡compPath _ _ _ + ∙∙ cong (emloop g ∙_) (isCommΩEM 0 q (sym (emloop g))) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ q) (rCancel (emloop g)) + ∙∙ sym (lUnit q))))) + ΩEM+1→EM-gen (suc n) = + Trunc.elim (λ _ → isOfHLevelΠ (4 + n) + λ _ → isOfHLevelSuc (3 + n) (hLevelEM _ (suc n))) + λ { north → ΩEM+1→EM (suc n) + ; south p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) + ∙∙ p + ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) + ; (merid a i) → help a i} + where + help : (a : EM-raw G (suc n)) + → PathP (λ i → Path (EM G (suc (suc n))) ∣ merid a i ∣ ∣ merid a i ∣ + → EM G (suc n)) + (ΩEM+1→EM (suc n)) + λ p → ΩEM+1→EM (suc n) (cong ∣_∣ₕ (merid ptEM-raw) + ∙∙ p + ∙∙ cong ∣_∣ₕ (sym (merid ptEM-raw))) + help a = + toPathP (funExt (λ p → + (transportRefl (ΩEM+1→EM (suc n) + (transport (λ i → Path (EM G (suc (suc n))) + ∣ merid a (~ i) ∣ ∣ merid a (~ i) ∣) p)) + ∙ cong (ΩEM+1→EM (suc n)) + (flipTransport + (((rUnit p + ∙ cong (p ∙_) (sym (lCancel _))) + ∙ isCommΩEM-base (suc n) ∣ south ∣ p _) + ∙∙ sym (∙assoc _ _ p) + ∙∙ cong₂ _∙_ (cong (cong ∣_∣ₕ) + (sym (cong sym (symDistr + (sym (merid a)) (merid ptEM-raw))))) + (isCommΩEM-base _ _ _ p) + ∙∙ sym (doubleCompPath≡compPath _ _ _) + ∙∙ λ j → transp (λ i → Path (EM G (suc (suc n))) + ∣ merid a (i ∨ ~ j) ∣ ∣ merid a (i ∨ ~ j) ∣) (~ j) + (cong ∣_∣ₕ (compPath-filler' + (sym (merid a)) (merid ptEM-raw) (~ j)) + ∙∙ p + ∙∙ cong ∣_∣ₕ (compPath-filler + (sym (merid ptEM-raw)) (merid a) (~ j)))))))) + + EM→ΩEM+1∘EM-raw→EM : (n : ℕ) (x : EM-raw G (suc n)) + → EM→ΩEM+1 (suc n) (EM-raw→EM _ _ x) ≡ cong ∣_∣ₕ (merid x ∙ sym (merid ptEM-raw)) + EM→ΩEM+1∘EM-raw→EM zero x = refl + EM→ΩEM+1∘EM-raw→EM (suc n) x = refl + + EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) + → EM G n → x ≡ x + EM→ΩEM+1-gen n x p = + sym (rUnitₖ (suc n) x) + ∙∙ cong (x +ₖ_) (EM→ΩEM+1 n p) + ∙∙ rUnitₖ (suc n) x + + ΩEM+1-gen→EM-0ₖ : (n : ℕ) (x : _) + → ΩEM+1→EM-gen n (0ₖ (suc n)) x + ≡ ΩEM+1→EM n x + ΩEM+1-gen→EM-0ₖ zero p = refl + ΩEM+1-gen→EM-0ₖ (suc n) p = refl + + EM→ΩEM+1-gen-0ₖ : (n : ℕ) (x : _) + → EM→ΩEM+1-gen n (0ₖ (suc n)) x + ≡ EM→ΩEM+1 n x + EM→ΩEM+1-gen-0ₖ zero x = sym (rUnit _) + ∙ λ j i → lUnitₖ 1 (EM→ΩEM+1 zero x i) j + EM→ΩEM+1-gen-0ₖ (suc n) x = sym (rUnit _) + ∙ λ j i → lUnitₖ (suc (suc n)) (EM→ΩEM+1 (suc n) x i) j + + EM→ΩEM+1→EM-gen : (n : ℕ) (x : EM G (suc n)) + → (y : EM G n) → ΩEM+1→EM-gen n x (EM→ΩEM+1-gen n x y) ≡ y + EM→ΩEM+1→EM-gen n = + EM-raw'-elim _ _ + (λ _ → isOfHLevelΠ (suc (suc n)) + (λ _ → isOfHLevelPath (suc (suc n)) + (hLevelEM _ n) _ _)) + (EM-raw'-trivElim _ n + (λ _ → isOfHLevelΠ (suc n) λ _ → hLevelEM _ n _ _) + λ y → cong (λ p → ΩEM+1→EM-gen n p + (EM→ΩEM+1-gen n p y)) + (EM-raw'→EM∙ G (suc n)) + ∙ (λ i → ΩEM+1-gen→EM-0ₖ n (EM→ΩEM+1-gen-0ₖ n y i) i) + ∙ Iso.leftInv (Iso-EM-ΩEM+1 n) y) + + ΩEM+1→EM→ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) + → (y : x ≡ x) → EM→ΩEM+1-gen n x (ΩEM+1→EM-gen n x y) ≡ y + ΩEM+1→EM→ΩEM+1-gen n = + EM-raw'-elim _ _ + (λ _ → isOfHLevelΠ (suc (suc n)) + (λ _ → isOfHLevelPath (suc (suc n)) + (hLevelEM _ (suc n) _ _) _ _)) + (EM-raw'-trivElim _ n + (λ _ → isOfHLevelΠ (suc n) + λ _ → hLevelEM _ (suc n) _ _ _ _) + (subst (λ p → (y : p ≡ p) + → EM→ΩEM+1-gen n p (ΩEM+1→EM-gen n p y) ≡ y) + (sym (EM-raw'→EM∙ _ (suc n))) + λ p → (λ i → EM→ΩEM+1-gen-0ₖ n (ΩEM+1-gen→EM-0ₖ n p i) i) + ∙ Iso.rightInv (Iso-EM-ΩEM+1 n) p)) + + Iso-EM-ΩEM+1-gen : (n : ℕ) (x : EM G (suc n)) + → Iso (EM G n) (x ≡ x) + Iso.fun (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1-gen n x + Iso.inv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM-gen n x + Iso.rightInv (Iso-EM-ΩEM+1-gen n x) = ΩEM+1→EM→ΩEM+1-gen n x + Iso.leftInv (Iso-EM-ΩEM+1-gen n x) = EM→ΩEM+1→EM-gen n x + + ΩEM+1→EM-gen-refl : (n : ℕ) (x : EM G (suc n)) + → ΩEM+1→EM-gen n x refl ≡ 0ₖ n + ΩEM+1→EM-gen-refl n = + EM-raw'-elim _ (suc n) + (λ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _) + (EM-raw'-trivElim _ n + (λ _ → hLevelEM _ n _ _) + (lem n)) + where + lem : (n : ℕ) → ΩEM+1→EM-gen n + (EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) refl + ≡ 0ₖ n + lem zero = ΩEM+1→EM-refl 0 + lem (suc n) = ΩEM+1→EM-refl (suc n) + + ΩEM+1→EM-gen-hom : (n : ℕ) (x : EM G (suc n)) (p q : x ≡ x) + → ΩEM+1→EM-gen n x (p ∙ q) ≡ ΩEM+1→EM-gen n x p +ₖ ΩEM+1→EM-gen n x q + ΩEM+1→EM-gen-hom n = + EM-raw'-elim _ (suc n) + (λ _ → isOfHLevelΠ2 (suc (suc n)) + (λ _ _ → isOfHLevelPath (suc (suc n)) (hLevelEM _ n) _ _)) + (EM-raw'-trivElim _ n + (λ _ → isOfHLevelΠ2 (suc n) (λ _ _ → hLevelEM _ n _ _)) + (lem n)) + where + lem : (n : ℕ) (p q : EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n))) + ≡ EM-raw'→EM G (suc n) (snd (EM-raw'∙ G (suc n)))) + → ΩEM+1→EM-gen n _ (p ∙ q) + ≡ ΩEM+1→EM-gen n _ p + +ₖ ΩEM+1→EM-gen n _ q + lem zero p q = ΩEM+1→EM-hom 0 p q + lem (suc n) p q = ΩEM+1→EM-hom (suc n) p q + +-- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for +-- the cup product +module _ where + open AbGroupStr renaming (_+_ to comp) + + isContr-↓∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ H n) + fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl + snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = + Σ≡Prop (λ x → is-set (snd H) _ _) + (funExt (raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) + fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl + fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = + Trunc.elim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) + (raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) + x i + snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + + isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) + → isContr ((EM-raw G (suc n) , ptEM-raw) →∙ EM∙ H n) + isContr-↓∙' zero = isContr-↓∙ zero + fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl + fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = + raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i + snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + + isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ EM∙ G n) + → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) + isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (EM∙ G (suc n))} {B = EM∙ G n} + (invEquiv (EM≃ΩEM+1 n)) + (ΩEM+1→EM-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ EM∙ G (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) + step₃ = isOfHLevelΩ→isOfHLevel m + λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) + (isHomogeneous→∙ (isHomogeneousEM (suc n)) f) + step₂ + + isOfHLevel↑∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n m → isOfHLevel (2 + n) (EM∙ G (suc m) →∙ EM∙ H (suc (n + m))) + isOfHLevel↑∙ zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙ m)) + isOfHLevel↑∙ (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + + isOfHLevel↑∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n m → isOfHLevel (2 + n) (EM-raw∙ G (suc m) →∙ EM∙ H (suc (n + m))) + isOfHLevel↑∙-lem zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) + isOfHLevel↑∙-lem (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙-lem n m) + + EM₁→∙Iso : {G : AbGroup ℓ} {H : AbGroup ℓ'} (m : ℕ) + → Iso (EM-raw'∙ G 1 →∙ EM∙ H (suc m)) (fst G → typ (Ω (EM∙ H (suc m)))) + Iso.fun (EM₁→∙Iso m) f g = sym (snd f) ∙∙ cong (fst f) (emloop-raw g) ∙∙ snd f + fst (Iso.inv (EM₁→∙Iso m) f) embase-raw = 0ₖ (suc m) + fst (Iso.inv (EM₁→∙Iso m) f) (emloop-raw g i) = f g i + snd (Iso.inv (EM₁→∙Iso m) f) = refl + Iso.rightInv (EM₁→∙Iso m) f = funExt λ x → sym (rUnit _) + Iso.leftInv (EM₁→∙Iso m) (f , p) = + →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ { embase-raw → sym p + ; (emloop-raw g i) j + → doubleCompPath-filler (sym p) (cong f (emloop-raw g)) p (~ j) i}) + + isOfHLevel↑∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n m → isOfHLevel (2 + n) (EM-raw'∙ G m →∙ EM∙ H (n + m)) + isOfHLevel↑∙' {H = H} zero zero = + isOfHLevelΣ 2 (isOfHLevelΠ 2 (λ _ → AbGroupStr.is-set (snd H))) + λ _ → isOfHLevelPath 2 (AbGroupStr.is-set (snd H)) _ _ + isOfHLevel↑∙' zero (suc zero) = + subst (isOfHLevel 2) (sym (isoToPath (EM₁→∙Iso 0))) + (isSetΠ λ _ → emsquash _ _) + isOfHLevel↑∙' zero (suc (suc m)) = isOfHLevel↑∙-lem zero (suc m) + isOfHLevel↑∙' {H = H} (suc n) zero = + isOfHLevelΣ (2 + suc n) (isOfHLevelΠ (2 + suc n) + (λ _ → subst (isOfHLevel (suc (suc (suc n)))) + (cong (EM H) (cong suc (+-comm 0 n))) + (hLevelEM H (suc n)))) + λ _ → isOfHLevelPath (suc (suc (suc n))) + (subst (isOfHLevel (suc (suc (suc n)))) + (cong (EM H) (cong suc (+-comm 0 n))) + (hLevelEM H (suc n))) _ _ + isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc zero) = + subst (isOfHLevel (2 + suc n)) (sym (isoToPath (EM₁→∙Iso (suc n))) + ∙ λ i → EM-raw'∙ G 1 →∙ EM∙ H (suc (+-comm 1 n i))) + (isOfHLevelΠ (2 + suc n) λ x → (isOfHLevelTrunc (4 + n) _ _)) + isOfHLevel↑∙' {G = G} {H = H} (suc n) (suc (suc m)) = + subst (isOfHLevel (2 + suc n)) + (λ i → (EM-raw'∙ G (suc (suc m)) + →∙ EM∙ H (suc (+-suc n (suc m) (~ i))))) + (isOfHLevel↑∙-lem (suc n) (suc m)) + + →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) + → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) + →∙EMPath {G = G} A n = + ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl + ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) + + private + contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) + → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) + fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl + snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = + →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) + (sym (funExt (help n f p))) + where + help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) + → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help zero f p = + raw-elim G zero + (λ _ → isOfHLevel↑∙ zero m _ _) p + help (suc n) f p = + Trunc.elim (λ _ → isOfHLevelPath (4 + n) + (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) + (λ i → suc (suc (+-comm (suc n) 1 i))) + (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) + (raw-elim G (suc n) + (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) + + contr∙-lem' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) + → isContr (EM∙ G (suc n) →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) + fst (contr∙-lem' n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl + snd (contr∙-lem' {G = G} {H = H} {L = L} n m) (f , p) = + →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) + (funExt λ x → sym (help' n f p x)) + where + help' : (n : ℕ) → (f : EM G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) + → f (snd (EM∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help' zero f p = + raw-elim _ zero (λ _ → isOfHLevel↑∙' zero (suc m) _ _) p + help' (suc n) f p = + Trunc.elim (λ _ → isOfHLevelPath (4 + n) + (subst2 (λ x y → isOfHLevel x (EM-raw'∙ H (suc m) →∙ EM∙ L y)) + (λ i → suc (suc (suc (+-comm n 1 i)))) + (cong suc (+-suc n m)) + (isOfHLevelPlus' {n = 1} (suc (suc (suc n))) + (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m)))) _ _) + (raw-elim _ (suc n) (λ _ → isOfHLevelPath' (2 + n) + (subst (λ y → isOfHLevel (suc (suc (suc n))) + (EM-raw'∙ H (suc m) →∙ EM∙ L y)) + (+-suc (suc n) m) + (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p) + + contr∙-lem'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) + → isContr (EM-raw'∙ G (suc n) + →∙ (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) + fst (contr∙-lem'' n m) = (λ _ → (λ _ → 0ₖ (suc (n + m))) , refl) , refl + snd (contr∙-lem'' {G = G} {H = H} {L = L} n m) (f , p) = + →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) + (funExt λ x → sym (help n f p x)) + where + help : (n : ℕ) → (f : EM-raw' G (suc n) → EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m))) + → f (snd (EM-raw'∙ G (suc n))) ≡ snd (EM-raw'∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help zero f p = + EM-raw'-trivElim G zero (λ _ → isOfHLevel↑∙' _ _ _ _) p + help (suc n) f p = + EM-raw'-trivElim _ _ (λ _ → isOfHLevelPath' (suc (suc n)) + (subst (λ y → isOfHLevel (suc (suc (suc n))) (EM-raw'∙ H (suc m) →∙ EM∙ L y)) + (cong suc (+-suc n m)) + (isOfHLevel↑∙' {G = H} {H = L} (suc n) (suc m))) _ _) p + + isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) + →∙ (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (l + n + m))) ∙)) + isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (EM∙ G (suc n) + →∙ (Ω (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (n + m))) ∙))) + h = + subst isProp + (λ i → EM∙ G (suc n) + →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) + isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst (isOfHLevel (2 + l)) + (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (EM∙ G (suc n) + →∙ (Ω (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → EM∙ G (suc n) + →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) + + isOfHLevel↑∙∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) + →∙ (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (l + n + m))) ∙)) + isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m zero = + isOfHLevelΩ→isOfHLevel 0 + λ f → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) + lem) + where + lem : isProp (EM∙ G (suc n) + →∙ (Ω (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (n + m))) ∙))) + lem = subst isProp + (λ i → EM∙ G (suc n) + →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem' n m)) + isOfHLevel↑∙∙' {G = G} {H = H} {L = L} n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst (isOfHLevel (2 + l)) + (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) + where + lem : isOfHLevel (2 + l) + (EM∙ G (suc n) + →∙ (Ω (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) + lem = + subst (isOfHLevel (2 + l)) + (λ i → EM∙ G (suc n) + →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙' n m l) + + isOfHLevel↑∙∙'' : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n m l → isOfHLevel (2 + l) (EM-raw'∙ G (suc n) + →∙ (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (l + n + m))) ∙)) + isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m zero = + isOfHLevelΩ→isOfHLevel 0 + λ f → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) + lem) + where + lem : isProp (EM-raw'∙ G (suc n) + →∙ (Ω (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (n + m))) ∙))) + lem = subst isProp + (λ i → EM-raw'∙ G (suc n) + →∙ (→∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem'' _ _)) + isOfHLevel↑∙∙'' {G = G} {H = H} {L = L} n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst (isOfHLevel (2 + l)) + (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) lem) + where + lem : isOfHLevel (2 + l) + (EM-raw'∙ G (suc n) + →∙ (Ω (EM-raw'∙ H (suc m) + →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) + lem = + subst (isOfHLevel (2 + l)) + (λ i → EM-raw'∙ G (suc n) + →∙ →∙EMPath {G = L} (EM-raw'∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙'' n m l) + +-- A homomorphism φ : G → H of AbGroups induces a homomorphism +-- φ' : K(G,n) → K(H,n) +inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupHom G' H' + → ∀ n + → EM-raw G' n → EM-raw H' n +inducedFun-EM-raw f = + elim+2 (fst f) + (EMrec _ emsquash embase + (λ g → emloop (fst f g)) + λ g h → compPathR→PathP (sym + (sym (lUnit _) + ∙∙ cong (_∙ (sym (emloop (fst f h)))) + (cong emloop (IsGroupHom.pres· (snd f) g h) + ∙ emloop-comp _ (fst f g) (fst f h)) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) + ∙∙ sym (rUnit _)))) + (λ n ind → λ { north → north + ; south → south + ; (merid a i) → merid (ind a) i} ) + +inducedFun-EM-raw-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM-raw G' n) + → inducedFun-EM-raw (idGroupHom {G = AbGroup→Group G'}) n x ≡ x +inducedFun-EM-raw-id zero x = refl +inducedFun-EM-raw-id (suc zero) = EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) + λ { embase-raw → refl ; (emloop-raw g i) → refl} +inducedFun-EM-raw-id (suc (suc n)) north = refl +inducedFun-EM-raw-id (suc (suc n)) south = refl +inducedFun-EM-raw-id (suc (suc n)) (merid a i) j = merid (inducedFun-EM-raw-id (suc n) a j) i + +inducedFun-EM-raw-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} + (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) + → (x : EM-raw G' n) → inducedFun-EM-raw (compGroupHom ϕ ψ) n x + ≡ inducedFun-EM-raw ψ n (inducedFun-EM-raw ϕ n x) +inducedFun-EM-raw-comp ϕ ψ zero x = refl +inducedFun-EM-raw-comp ϕ ψ (suc zero) = + EM-raw'-elim _ 1 (λ _ → hLevelEM _ 1 _ _) + λ { embase-raw → refl ; (emloop-raw g i) → refl} +inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) north = refl +inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) south = refl +inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) (merid a i) j = + merid (inducedFun-EM-raw-comp ϕ ψ (suc n) a j) i + +inducedFun-EM : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupHom G' H' + → ∀ n + → EM G' n → EM H' n +inducedFun-EM f zero = inducedFun-EM-raw f zero +inducedFun-EM f (suc zero) = inducedFun-EM-raw f (suc zero) +inducedFun-EM f (suc (suc n)) = Trunc.map (inducedFun-EM-raw f (2 + n)) + +EM-raw→EM-funct : {G : AbGroup ℓ} {H : AbGroup ℓ'} + (n : ℕ) (ψ : AbGroupHom G H) (y : EM-raw G n) + → EM-raw→EM _ _ (inducedFun-EM-raw ψ n y) + ≡ inducedFun-EM ψ n (EM-raw→EM _ _ y) +EM-raw→EM-funct zero ψ y = refl +EM-raw→EM-funct (suc zero) ψ y = refl +EM-raw→EM-funct (suc (suc n)) ψ y = refl + +inducedFun-EM-id : {G' : AbGroup ℓ} (n : ℕ) (x : EM G' n) + → inducedFun-EM (idGroupHom {G = AbGroup→Group G'}) n x ≡ x +inducedFun-EM-id zero x = refl +inducedFun-EM-id (suc zero) = inducedFun-EM-raw-id (suc zero) +inducedFun-EM-id (suc (suc n)) = + Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) + λ x → cong ∣_∣ₕ (inducedFun-EM-raw-id _ x) + +inducedFun-EM-comp : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {L' : AbGroup ℓ''} + (ϕ : AbGroupHom G' H') (ψ : AbGroupHom H' L') (n : ℕ) + → (x : EM G' n) → inducedFun-EM (compGroupHom ϕ ψ) n x + ≡ inducedFun-EM ψ n (inducedFun-EM ϕ n x) +inducedFun-EM-comp ϕ ψ zero x = refl +inducedFun-EM-comp ϕ ψ (suc zero) = inducedFun-EM-raw-comp ϕ ψ (suc zero) +inducedFun-EM-comp ϕ ψ (suc (suc n)) = + Trunc.elim (λ _ → isOfHLevelPath (4 + n) (hLevelEM _ (suc (suc n))) _ _) + λ x → cong ∣_∣ₕ (inducedFun-EM-raw-comp ϕ ψ (suc (suc n)) x) + +inducedFun-EM0ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} {ϕ : AbGroupHom G' H'} (n : ℕ) + → inducedFun-EM ϕ n (0ₖ n) ≡ 0ₖ n +inducedFun-EM0ₖ {ϕ = ϕ} zero = IsGroupHom.pres1 (snd ϕ) +inducedFun-EM0ₖ (suc zero) = refl +inducedFun-EM0ₖ (suc (suc n)) = refl + +inducedFun-EM-pres+ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + (ϕ : AbGroupHom G' H') (n : ℕ) (x y : EM G' n) + → inducedFun-EM ϕ n (x +ₖ y) ≡ inducedFun-EM ϕ n x +ₖ inducedFun-EM ϕ n y +inducedFun-EM-pres+ₖ ϕ zero x y = IsGroupHom.pres· (snd ϕ) x y +inducedFun-EM-pres+ₖ {G' = G'} {H' = H'} ϕ (suc n) = + EM-elim2 (suc n) (λ _ _ → isOfHLevelPath (2 + suc n) (hLevelEM _ (suc n)) _ _) + (wedgeConEM.fun _ _ n n + (λ _ _ → isOfHLevelPath' (suc n + suc n) + (subst (λ m → isOfHLevel (suc (suc m)) (EM H' (suc n))) + (sym (+-suc n n)) + (isOfHLevelPlus' {n = n} (3 + n) + (hLevelEM _ (suc n)))) + _ _) + (l n) + (r n) + (l≡r n)) + where + lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) → EM-raw→EM G (suc n) ptEM-raw ≡ 0ₖ _ + lem zero = refl + lem (suc n) = refl + + l : (n : ℕ) (y : EM-raw G' (suc n)) + → inducedFun-EM ϕ (suc n) ((EM-raw→EM G' (suc n) ptEM-raw) + +ₖ EM-raw→EM G' (suc n) y) + ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) + +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y) + l n y = (cong (inducedFun-EM ϕ (suc n)) + (cong (_+ₖ EM-raw→EM G' (suc n) y) (lem n) + ∙ lUnitₖ (suc n) (EM-raw→EM G' (suc n) y))) + ∙∙ sym (lUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) + ∙∙ cong (_+ₖ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) y))) + (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) + ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) + + r : (n : ℕ) (x : EM-raw G' (suc n)) + → inducedFun-EM ϕ (suc n) + (EM-raw→EM G' (suc n) x +ₖ EM-raw→EM G' (suc n) ptEM-raw) + ≡ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) + +ₖ inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) ptEM-raw) + r n x = cong (inducedFun-EM ϕ (suc n)) + (cong (EM-raw→EM G' (suc n) x +ₖ_) (lem n) + ∙ rUnitₖ (suc n) (EM-raw→EM G' (suc n) x)) + ∙∙ sym (rUnitₖ _ (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x))) + ∙∙ cong (inducedFun-EM ϕ (suc n) (EM-raw→EM G' (suc n) x) +ₖ_) + (sym (inducedFun-EM0ₖ {ϕ = ϕ} (suc n)) + ∙ cong (inducedFun-EM ϕ (suc n)) (sym (lem n))) + + l≡r : (n : ℕ) → l n ptEM-raw ≡ r n ptEM-raw + l≡r zero = refl + l≡r (suc n) = refl + +inducedFun-EM-pres-ₖ : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + (ϕ : AbGroupHom G' H') (n : ℕ) (x : EM G' n) + → inducedFun-EM ϕ n (-ₖ x) ≡ -ₖ (inducedFun-EM ϕ n x) +inducedFun-EM-pres-ₖ ϕ n = + morphLemmas.distrMinus _+ₖ_ _+ₖ_ + (inducedFun-EM ϕ n) (inducedFun-EM-pres+ₖ ϕ n) (0ₖ n) (0ₖ n) (-ₖ_) (-ₖ_) + (lUnitₖ n) (rUnitₖ n) + (lCancelₖ n) (rCancelₖ n) + (assocₖ n) (inducedFun-EM0ₖ n) + +EMFun-EM→ΩEM+1 : {G : AbGroup ℓ} {H : AbGroup ℓ'} + {ϕ : AbGroupHom G H} (n : ℕ) (x : EM G n) + → PathP (λ i → inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i) + ≡ inducedFun-EM0ₖ {ϕ = ϕ} (suc n) (~ i)) + (EM→ΩEM+1 n (inducedFun-EM ϕ n x)) + (cong (inducedFun-EM ϕ (suc n)) (EM→ΩEM+1 n x)) +EMFun-EM→ΩEM+1 {ϕ = ϕ} zero x = refl +EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc zero) x = + cong-∙ ∣_∣ₕ (merid (inducedFun-EM ϕ (suc zero) x)) + (sym (merid embase)) + ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc zero))) + (cong ∣_∣ₕ (merid x)) (cong ∣_∣ₕ (sym (merid embase)))) + ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc zero)))) + (sym (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase)))) +EMFun-EM→ΩEM+1 {ϕ = ϕ} (suc (suc n)) = + Trunc.elim (λ _ → isOfHLevelPath (4 + n) + (isOfHLevelTrunc (5 + n) _ _) _ _) + λ a → cong-∙ ∣_∣ₕ (merid (inducedFun-EM-raw ϕ (2 + n) a)) + (sym (merid north)) + ∙∙ sym (cong-∙ (inducedFun-EM ϕ (suc (suc (suc n)))) + (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (sym (merid north)))) + ∙∙ cong (cong (inducedFun-EM ϕ (suc (suc (suc n))))) + (sym (cong-∙ ∣_∣ₕ (merid a) (sym (merid north)))) + + +inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupEquiv G' H' + → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +Iso.rightInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) + (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) + h = elim+2 + (secEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + (sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) + (sym (lUnit _)) + ∙ rCancel _)))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} +Iso.leftInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) + (Iso.inv (inducedFun-EM-rawIso e n)) + h = elim+2 + (retEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} + +module _ {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) where + Iso→EMIso : ∀ n → Iso (EM G n) (EM H n) + Iso.fun (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom e) n + Iso.inv (Iso→EMIso n) = inducedFun-EM (GroupEquiv→GroupHom (invGroupEquiv e)) n + Iso.rightInv (Iso→EMIso zero) = Iso.rightInv (inducedFun-EM-rawIso e zero) + Iso.rightInv (Iso→EMIso (suc zero)) = + Iso.rightInv (inducedFun-EM-rawIso e (suc zero)) + Iso.rightInv (Iso→EMIso (suc (suc n))) = + Iso.rightInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) + Iso.leftInv (Iso→EMIso zero) = + Iso.leftInv (inducedFun-EM-rawIso e zero) + Iso.leftInv (Iso→EMIso (suc zero)) = + Iso.leftInv (inducedFun-EM-rawIso e (suc zero)) + Iso.leftInv (Iso→EMIso (suc (suc n))) = + Iso.leftInv (mapCompIso (inducedFun-EM-rawIso e (suc (suc n)))) + + Iso→EMIso∙ : ∀ n → Iso.fun (Iso→EMIso n) (EM∙ G n .snd) ≡ EM∙ H n .snd + Iso→EMIso∙ zero = IsGroupHom.pres1 (e .snd) + Iso→EMIso∙ (suc zero) = refl + Iso→EMIso∙ (suc (suc n)) = refl + + Iso→EMIso⁻∙ : ∀ n → Iso.inv (Iso→EMIso n) (EM∙ H n .snd) ≡ EM∙ G n .snd + Iso→EMIso⁻∙ zero = IsGroupHom.pres1 (invGroupEquiv e .snd) + Iso→EMIso⁻∙ (suc zero) = refl + Iso→EMIso⁻∙ (suc (suc n)) = refl + +Iso→EMIsoInv : {G : AbGroup ℓ} {H : AbGroup ℓ'} (e : AbGroupEquiv G H) + → ∀ n → Iso.inv (Iso→EMIso e n) ≡ Iso.fun (Iso→EMIso (invGroupEquiv e) n) +Iso→EMIsoInv e zero = refl +Iso→EMIsoInv e (suc zero) = refl +Iso→EMIsoInv e (suc (suc n)) = refl + +EM⊗-commIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +EM⊗-commIso {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) From 276623f26be68b3fa990aabaf9858e2a169aec53 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Tue, 2 Apr 2024 04:05:33 +0200 Subject: [PATCH 13/20] cleanup --- Cubical/Tactics/PathSolver/Example.agda | 2 -- 1 file changed, 2 deletions(-) diff --git a/Cubical/Tactics/PathSolver/Example.agda b/Cubical/Tactics/PathSolver/Example.agda index 3cdb601bd2..8b8862dd2f 100644 --- a/Cubical/Tactics/PathSolver/Example.agda +++ b/Cubical/Tactics/PathSolver/Example.agda @@ -81,7 +81,5 @@ module Examples4 (f : B → A) (g : C → B) module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where - open import Cubical.Foundations.Equiv - SqCompEquiv : (Square p p q q) ≃ (p ∙ q ≡ q ∙ p) SqCompEquiv = 2-cylinder-from-square.Sq≃Sq' refl ≡! From d8dc7a568775f381a6e89e716f5537b15a72c9b9 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 15 Apr 2024 02:35:25 +0200 Subject: [PATCH 14/20] got rid of wraping of GroupMorphism as WildCatFunctor --- Cubical/Tactics/GroupSolver/Example.agda | 21 ++- Cubical/Tactics/GroupSolver/Solver.agda | 49 ++++--- Cubical/Tactics/GroupoidSolver/Example.agda | 134 ++++++++++-------- Cubical/Tactics/GroupoidSolver/Solver.agda | 31 +++- Cubical/Tactics/Reflection.agda | 5 + Cubical/Tactics/WildCatSolver/Example.agda | 17 ++- Cubical/Tactics/WildCatSolver/Reflection.agda | 20 +++ Cubical/Tactics/WildCatSolver/Solver.agda | 36 ++++- Cubical/Tactics/WildCatSolver/Solvers.agda | 53 ++++--- 9 files changed, 241 insertions(+), 125 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index fdf4dd9f7b..dc9b300b68 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -5,18 +5,19 @@ module Cubical.Tactics.GroupSolver.Example where open import Cubical.Foundations.Prelude open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties + open import Cubical.Tactics.GroupSolver.Solver open import Cubical.Data.List -open import Cubical.WildCat.Functor - private variable ℓ : Level module example (G G* G○ : Group ℓ) - (F* : GroupHom' G* G) - (F○ : GroupHom' G○ G*) + (F* : GroupHom G* G) + (F○ : GroupHom G○ G*) where @@ -25,8 +26,6 @@ module example (G G* G○ : Group ℓ) open GroupStr (snd G) - open WildFunctor - module * where open GroupStr (snd G*) public @@ -46,16 +45,16 @@ module example (G G* G○ : Group ℓ) pA≡pB = solveGroup (G ∷ []) - module T2 p p' q r s t u where lhs rhs : fst G - lhs = p · (p · inv p) · inv p · (p' · inv p') · (p · p') · (inv p' · (F* ⟪ (((*.inv q) *.· r) *.· F○ ⟪ ○.inv t ⟫ *.· - (*.inv (F○ ⟪ s ○.· s ⟫) *.· F○ ⟪ u ⟫ )) ⟫ )) + lhs = p · (p · inv p) · inv p · (p' · inv p') · (p · p') · + (inv p' · (fst F* (((*.inv q) *.· r) *.· fst F○ (○.inv t) *.· + (*.inv (fst F○ ( s ○.· s )) *.· fst F○ u )) )) - rhs = inv (F* ⟪ q ⟫ · inv p) · (F* ⟪ r ⟫) · - (comp-WildFunctor F○ F*) ⟪ ○.inv (s ○.· s ○.· t) ○.· u ⟫ + rhs = inv (fst F* q · inv p) · (fst F* r) · + fst (compGroupHom F○ F*) (○.inv (s ○.· s ○.· t) ○.· u ) lhs≡rhs : lhs ≡ rhs diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index 602abeb41b..01c2ca3121 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -9,18 +9,22 @@ open import Cubical.Foundations.Function open import Cubical.Data.Unit open import Cubical.Data.List open import Cubical.Data.Maybe +open import Cubical.Data.Sigma open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R open import Cubical.WildCat.Base -open import Cubical.Tactics.WildCatSolver.Solvers + + + +open import Cubical.Tactics.WildCatSolver.Reflection open import Cubical.WildCat.Functor open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Tactics.Reflection open import Cubical.Tactics.WildCatSolver.Solvers @@ -40,38 +44,33 @@ module _ {ℓ} where wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f +module Group-Solver ℓ where - GroupHom' : (G H : Group ℓ) → Type ℓ - - GroupHom' G H = WildFunctor - (WildGroupoid.wildCat (Group→WildGroupoid G)) - (WildGroupoid.wildCat (Group→WildGroupoid H)) + mbGroupHomApp : R.Term → Maybe (R.Term × R.Term) + mbGroupHomApp (R.def (quote fst) t) = match2Vargs' t + mbGroupHomApp _ = nothing - IsoGroupHom' : ∀ {G H} → Iso (GroupHom' G H) (GroupHom G H) - Iso.fun IsoGroupHom' wf = _ , makeIsGroupHom (WildFunctor.F-seq wf) - WildFunctor.F-ob (Iso.inv IsoGroupHom' _) = λ _ → tt - WildFunctor.F-hom (Iso.inv IsoGroupHom' (f , _)) = f - WildFunctor.F-id (Iso.inv IsoGroupHom' (_ , gh)) = IsGroupHom.pres1 gh - WildFunctor.F-seq (Iso.inv IsoGroupHom' (_ , gh)) = IsGroupHom.pres· gh - Iso.rightInv IsoGroupHom' _ = GroupHom≡ refl - WildFunctor.F-ob (Iso.leftInv IsoGroupHom' _ i) = λ _ → tt - WildFunctor.F-hom (Iso.leftInv IsoGroupHom' wf i) = WildFunctor.F-hom wf - WildFunctor.F-id (Iso.leftInv (IsoGroupHom' {G = G} {H = H}) wf i) = - IsGroup.is-set (GroupStr.isGroup (snd H)) - (WildFunctor.F-hom wf (GroupStr.1g (snd G))) - (GroupStr.1g (snd H)) - (hom1g (G .snd) (WildFunctor.F-hom wf) (H .snd) - (WildFunctor.F-seq wf)) - (WildFunctor.F-id wf) i - WildFunctor.F-seq (Iso.leftInv IsoGroupHom' wf i) = λ f g → WildFunctor.F-seq wf f g + groupHomTrmSrc : R.Term → R.TC R.Term + groupHomTrmSrc (R.def (quote Σ) t) = do + (fun , isGrpHomTrm) ← match2Vargs t + funDom ← matchPiDom fun + unFst funDom + groupHomTrmSrc t = R.typeError $ (R.strErr "groupHomTrmSrc fail: ") ∷ (getConTail t) -module Group-Solver ℓ where GroupWS : WildCatInstance ℓ-zero ℓ WildCatInstance.wildStr GroupWS = Group ℓ WildCatInstance.toWildCat GroupWS = WildGroupoid.wildCat ∘ Group→WildGroupoid WildCatInstance.mbIsWildGroupoid GroupWS = just (WildGroupoid.isWildGroupoid ∘ Group→WildGroupoid) + WildCatInstance.wildStrMor GroupWS = GroupHom + WildCatInstance.toWildFunctor GroupWS _ _ (f , isGHom) = + record { F-ob = _ ; F-hom = f ; F-id = pres1 ; F-seq = pres· } + where open IsGroupHom isGHom + WildCatInstance.mbFunctorApp GroupWS = mbGroupHomApp + WildCatInstance.F-ty-extractSrc GroupWS = + groupHomTrmSrc >=& λ x → + R.def (quote WildGroupoid.wildCat) v[ R.def (quote Group→WildGroupoid) (v[ x ])] private module GRP-WS = WildCatInstance GroupWS diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 182996eb9f..4b07ce4b9d 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -9,105 +9,117 @@ open import Cubical.Tactics.GroupoidSolver.Solver open import Cubical.Data.List open import Cubical.Data.Nat open import Cubical.Categories.Category -open import Cubical.WildCat.Functor private variable ℓ ℓ' : Level -module example (WG WG* : WildGroupoid ℓ ℓ') - (F : WildFunctor - (WildGroupoid.wildCat WG*) - (WildGroupoid.wildCat WG)) - where - open WildGroupoid-Solver ℓ ℓ' +module WildGroupoidExamples where + open import Cubical.WildCat.Functor - open WildGroupoid WG - open WildCat wildCat - open WildFunctor + module example (WG WG* : WildGroupoid ℓ ℓ') + (F : WildFunctor + (WildGroupoid.wildCat WG*) + (WildGroupoid.wildCat WG)) + where + open WildGroupoid-Solver ℓ ℓ' - module * where - open WildCat (WildGroupoid.wildCat WG*) public - open WildGroupoid WG* public + open WildGroupoid WG + open WildCat wildCat - module T1 x y z w - (p p' : Hom[ x , y ]) - (q : Hom[ y , z ]) - (r : Hom[ z , w ]) where + open WildFunctor - pA pB : Hom[ x , w ] - pA = ((((((id ⋆ p') ⋆ q) ⋆ ((inv q) ⋆ id)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r - pB = p ⋆ (q ⋆ r) + module * where + open WildCat (WildGroupoid.wildCat WG*) public + open WildGroupoid WG* public - pA≡pB : pA ≡ pB - pA≡pB = solveWildGroupoid (WG ∷ WG* ∷ []) + module T1 x y z w + (p p' : Hom[ x , y ]) + (q : Hom[ y , z ]) + (r : Hom[ z , w ]) where - module T2 x y z w - (p : Hom[ x , F ⟅ y ⟆ ]) - (p' : Hom[ F ⟅ y ⟆ , x ]) - (q : *.Hom[ z , y ]) - (r : *.Hom[ z , w ]) where + pA pB : Hom[ x , w ] + pA = ((((((id ⋆ p') ⋆ q) ⋆ ((inv q) ⋆ id)) ⋆ (inv p')) ⋆ p) ⋆ q) ⋆ r + pB = p ⋆ (q ⋆ r) + pA≡pB : pA ≡ pB + pA≡pB = solveWildGroupoid (WG ∷ WG* ∷ []) - lhs rhs : Hom[ x , F ⟅ w ⟆ ] - lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F ⟪ (*.inv q) *.⋆ r ⟫)) - rhs = inv (F ⟪ q ⟫ ⋆ inv p) ⋆ F ⟪ r ⟫ - lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (WG ∷ WG* ∷ []) + module T2 x y z w + (p : Hom[ x , F ⟅ y ⟆ ]) + (p' : Hom[ F ⟅ y ⟆ , x ]) + (q : *.Hom[ z , y ]) + (r : *.Hom[ z , w ]) where - module T3 (obs : ℕ → ob) - (homs : ∀ x y → ℕ → Hom[ obs x , obs y ]) - where + lhs rhs : Hom[ x , F ⟅ w ⟆ ] + lhs = (p ⋆ p') ⋆ (inv p' ⋆ (F ⟪ (*.inv q) *.⋆ r ⟫)) + rhs = inv (F ⟪ q ⟫ ⋆ inv p) ⋆ F ⟪ r ⟫ - lhs rhs : Hom[ obs 1 , obs 5 ] - lhs = homs _ 2 1 ⋆ (homs _ _ 2 ⋆ (homs 3 5 2 ⋆ (homs _ 6 3 ⋆ inv (homs _ 6 3)))) - rhs = ((homs _ 2 1 ⋆ homs _ _ 2) ⋆ id) ⋆ homs 3 5 2 + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (WG ∷ WG* ∷ []) - lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (WG ∷ []) + module T3 (obs : ℕ → ob) + (homs : ∀ x y → ℕ → Hom[ obs x , obs y ]) + where -module exampleGpd (G G* : GroupoidCat ℓ ℓ') - (F : WildFunctor - (WildGroupoid.wildCat (Groupoid-Solver.Groupoid→WildGroupoid _ _ G*)) - (WildGroupoid.wildCat (Groupoid-Solver.Groupoid→WildGroupoid _ _ G))) - where + lhs rhs : Hom[ obs 1 , obs 5 ] + lhs = homs _ 2 1 ⋆ (homs _ _ 2 ⋆ (homs 3 5 2 ⋆ (homs _ 6 3 ⋆ inv (homs _ 6 3)))) + rhs = ((homs _ 2 1 ⋆ homs _ _ 2) ⋆ id) ⋆ homs 3 5 2 - open Groupoid-Solver ℓ ℓ' + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (WG ∷ []) - module * where - open GroupoidCat G* public - open Category category public +module GroupoidExamples where - open GroupoidCat G - open Category category + open import Cubical.Categories.Functor - module T1 x y z w - (p : Hom[ x , F ⟅ y ⟆ ]) - (p' : Hom[ F ⟅ y ⟆ , x ]) - (q : *.Hom[ z , y ]) - (r : *.Hom[ z , w ]) where + module exampleGpd (G G* : GroupoidCat ℓ ℓ') + (F : Functor + (GroupoidCat.category G*) + (GroupoidCat.category G)) + where - lhs rhs : Hom[ x , F ⟅ w ⟆ ] - lhs = (p ⋆ p') ⋆ (p' ⁻¹ ⋆ (F ⟪ (q *.⁻¹) *.⋆ r ⟫)) - rhs = (F ⟪ q ⟫ ⋆ p ⁻¹) ⁻¹ ⋆ F ⟪ r ⟫ + open Groupoid-Solver ℓ ℓ' - lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (G ∷ G* ∷ []) + + module * where + open GroupoidCat G* public + open Category category public + + open GroupoidCat G + open Category category + + + + module T1 x y z w + (p : Hom[ x , F ⟅ y ⟆ ]) + (p' : Hom[ F ⟅ y ⟆ , x ]) + (q : *.Hom[ z , y ]) + (r : *.Hom[ z , w ]) where + + + lhs rhs : Hom[ x , F ⟅ w ⟆ ] + lhs = (p ⋆ p') ⋆ (p' ⁻¹ ⋆ (F ⟪ (q *.⁻¹) *.⋆ r ⟫)) + rhs = (F ⟪ q ⟫ ⋆ p ⁻¹) ⁻¹ ⋆ F ⟪ r ⟫ + + + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveWildGroupoid (G ∷ G* ∷ []) diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index a5cd64d754..d174eae868 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -7,23 +7,38 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Data.Unit open import Cubical.Data.List +open import Cubical.Data.Sigma open import Cubical.Data.Maybe open import Cubical.Reflection.Base import Agda.Builtin.Reflection as R open import Cubical.WildCat.Base + open import Cubical.Categories.Category + +open import Cubical.Tactics.Reflection open import Cubical.Tactics.WildCatSolver.Solvers +open import Cubical.Tactics.WildCatSolver.Reflection open import Cubical.Tactics.WildCatSolver.Solver module WildGroupoid-Solver ℓ ℓ' where + open import Cubical.WildCat.Functor + + mbWildFunctorApp : R.Term → Maybe (R.Term × R.Term) + mbWildFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t + mbWildFunctorApp _ = nothing + + GroupoidWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr GroupoidWS = WildGroupoid ℓ ℓ' WildCatInstance.toWildCat GroupoidWS = WildGroupoid.wildCat WildCatInstance.mbIsWildGroupoid GroupoidWS = just WildGroupoid.isWildGroupoid - + WildCatInstance.wildStrMor GroupoidWS _ _ = WildFunctor _ _ + WildCatInstance.toWildFunctor GroupoidWS _ _ f = f + WildCatInstance.mbFunctorApp GroupoidWS = mbWildFunctorApp + WildCatInstance.F-ty-extractSrc GroupoidWS = extraxtWildFunSrc private module WGPD-WS = WildCatInstance GroupoidWS @@ -32,6 +47,11 @@ module WildGroupoid-Solver ℓ ℓ' where solveWildGroupoid = WGPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) module Groupoid-Solver ℓ ℓ' where + open import Cubical.Categories.Functor + + mbFunctorApp : R.Term → Maybe (R.Term × R.Term) + mbFunctorApp (R.def (quote Functor.F-hom) t) = matchFunctorAppArgs t + mbFunctorApp _ = nothing Groupoid→WildGroupoid : GroupoidCat ℓ ℓ' → WildGroupoid ℓ ℓ' @@ -52,7 +72,14 @@ module Groupoid-Solver ℓ ℓ' where WildCatInstance.toWildCat GroupoidWS = WildGroupoid.wildCat ∘ Groupoid→WildGroupoid WildCatInstance.mbIsWildGroupoid GroupoidWS = just (WildGroupoid.isWildGroupoid ∘ Groupoid→WildGroupoid) - + WildCatInstance.wildStrMor GroupoidWS x y = Functor (GroupoidCat.category x) (GroupoidCat.category y) + WildCatInstance.toWildFunctor GroupoidWS x y f = + record { F-ob = F-ob ; F-hom = F-hom ; F-id = F-id ; F-seq = F-seq } + where open Functor f + WildCatInstance.mbFunctorApp GroupoidWS = mbFunctorApp + WildCatInstance.F-ty-extractSrc GroupoidWS = + extraxtWildFunSrc >=& λ x → + R.def (quote Cat-Solver.Cat→WildCat) (R.unknown v∷ R.unknown v∷ v[ x ]) private module GPD-WS = WildCatInstance GroupoidWS diff --git a/Cubical/Tactics/Reflection.agda b/Cubical/Tactics/Reflection.agda index 0e08a7a3f0..929e4e6490 100644 --- a/Cubical/Tactics/Reflection.agda +++ b/Cubical/Tactics/Reflection.agda @@ -33,6 +33,11 @@ _>=>_ : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} → (A → TC B) → (B → TC C) → A → TC C (x >=> x₁) x₂ = x x₂ >>= x₁ +_>=&_ : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ}{B : Type ℓ'}{C : Type ℓ''} → + (A → TC B) → (B → C) → A → TC C +(x >=& x₁) x₂ = x x₂ >>= (λ u → pure (x₁ u)) + + wait-for-term-args : List (Arg Term) → TC (List (Arg Term)) wait-for-term-clauses : List (Clause) → TC (List Clause) wait-for-term-clause : Clause → TC Clause diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index 1d725db2d3..22a9c00aeb 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -7,7 +7,6 @@ open import Cubical.Foundations.Prelude open import Cubical.WildCat.Base open import Cubical.Tactics.WildCatSolver.Solver open import Cubical.Data.List -open import Cubical.WildCat.Functor private variable @@ -15,6 +14,7 @@ private module exampleWC where + open import Cubical.WildCat.Functor module _ (WC WC* : WildCat ℓ ℓ') (F : WildFunctor WC* WC) where @@ -45,10 +45,12 @@ module exampleWC where module exampleC ℓ ℓ' where open import Cubical.Categories.Category + open import Cubical.Categories.Functor + open Cat-Solver ℓ ℓ' - module _ (C C* : Category ℓ ℓ') (F : Functor' C* C) where + module _ (C C* : Category ℓ ℓ') (F : Functor C* C) where open Category C @@ -80,6 +82,17 @@ module exampleC ℓ ℓ' where (s : Hom[ F ⟅ w ⟆ , v ]) where + + + + -- pA pB : Hom[ F ⟅ y ⟆ , F ⟅ w ⟆ ] + -- pA = F ⟪ q *.⋆ r ⟫ + -- pB = F ⟪ q ⟫ ⋆ F ⟪ r ⟫ + + -- pA=pB : pA ≡ pB + -- pA=pB = solveCat (C ∷ C* ∷ []) + + pA pB pC : Hom[ x , v ] pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) diff --git a/Cubical/Tactics/WildCatSolver/Reflection.agda b/Cubical/Tactics/WildCatSolver/Reflection.agda index 10cf2ca65f..63b97d7157 100644 --- a/Cubical/Tactics/WildCatSolver/Reflection.agda +++ b/Cubical/Tactics/WildCatSolver/Reflection.agda @@ -118,12 +118,23 @@ matchVarg (harg _ ∷ xs) = matchVarg xs matchVarg (varg t ∷ []) = R.returnTC t matchVarg _ = R.typeError [ R.strErr "matchV fail" ] +matchFirstVarg : List (R.Arg R.Term) → R.TC R.Term +matchFirstVarg (harg _ ∷ xs) = matchFirstVarg xs +matchFirstVarg (varg t ∷ _) = R.returnTC t +matchFirstVarg _ = R.typeError [ R.strErr "matchV fail" ] + + match2Vargs : List (R.Arg R.Term) → R.TC (R.Term × R.Term) match2Vargs (harg _ ∷ xs) = match2Vargs xs match2Vargs (varg t1 ∷ varg t2 ∷ []) = R.returnTC (t1 , t2) match2Vargs _ = R.typeError [] +match2Vargs' : List (R.Arg R.Term) → Maybe (R.Term × R.Term) +match2Vargs' (harg _ ∷ xs) = match2Vargs' xs +match2Vargs' (varg t1 ∷ varg t2 ∷ []) = just (t1 , t2) +match2Vargs' _ = nothing + matchFunctorAppArgs : List (R.Arg R.Term) → Maybe (R.Term × R.Term) matchFunctorAppArgs (harg _ ∷ xs) = matchFunctorAppArgs xs matchFunctorAppArgs (varg t1 ∷ harg _ ∷ harg _ ∷ varg t2 ∷ []) = just (t1 , t2) @@ -256,3 +267,12 @@ foldPathTerms (just x ∷ xs) = symPathTerms : List (Maybe R.Term) → List (Maybe R.Term) symPathTerms = map (map-Maybe (R.def (quote sym) ∘ v[_])) ∘ rev + + +matchPiDom : R.Term → R.TC R.Term +matchPiDom (R.pi (varg d) _) = pure d +matchPiDom t = R.typeError ("matchPiDom fail" ∷ₑ [ t ]ₑ ) + +unFst : R.Term → R.TC R.Term +unFst (R.def (quote fst) t) = matchVarg t +unFst t = R.typeError ("unFst fail" ∷ₑ [ t ]ₑ ) diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index 5e7b3f84e5..d56c064e9f 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -3,8 +3,10 @@ module Cubical.Tactics.WildCatSolver.Solver where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Data.Unit open import Cubical.Data.List +open import Cubical.Data.Sigma open import Cubical.Data.Maybe open import Cubical.Reflection.Base @@ -12,15 +14,29 @@ import Agda.Builtin.Reflection as R open import Cubical.WildCat.Base open import Cubical.WildCat.Functor +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.WildCatSolver.Reflection open import Cubical.Tactics.WildCatSolver.Solvers open import Cubical.Categories.Category +open import Cubical.Categories.Functor + module WildCat-Solver ℓ ℓ' where + + mbWildFunctorApp : R.Term → Maybe (R.Term × R.Term) + mbWildFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t + mbWildFunctorApp _ = nothing + + WildCatWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr (WildCatWS) = WildCat ℓ ℓ' WildCatInstance.toWildCat WildCatWS x = x WildCatInstance.mbIsWildGroupoid WildCatWS = nothing + WildCatInstance.wildStrMor WildCatWS x y = WildFunctor x y + WildCatInstance.toWildFunctor WildCatWS _ _ f = f + WildCatInstance.mbFunctorApp WildCatWS = mbWildFunctorApp + WildCatInstance.F-ty-extractSrc WildCatWS = extraxtWildFunSrc private module WC-WS = WildCatInstance WildCatWS @@ -32,6 +48,13 @@ module WildCat-Solver ℓ ℓ' where module Cat-Solver ℓ ℓ' where + + + mbFunctorApp : R.Term → Maybe (R.Term × R.Term) + mbFunctorApp (R.def (quote Functor.F-hom) t) = matchFunctorAppArgs t + mbFunctorApp _ = nothing + + Cat→WildCat : Category ℓ ℓ' → WildCat ℓ ℓ' WildCat.ob (Cat→WildCat x) = Category.ob x WildCat.Hom[_,_] (Cat→WildCat x) = Category.Hom[_,_] x @@ -41,14 +64,19 @@ module Cat-Solver ℓ ℓ' where WildCat.⋆IdR (Cat→WildCat x) = Category.⋆IdR x WildCat.⋆Assoc (Cat→WildCat x) = Category.⋆Assoc x + CatWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr CatWS = Category ℓ ℓ' WildCatInstance.toWildCat CatWS = Cat→WildCat WildCatInstance.mbIsWildGroupoid CatWS = nothing - - Functor' : Category ℓ ℓ' → Category ℓ ℓ' → Type (ℓ-max ℓ ℓ') - Functor' x x₁ = WildFunctor (Cat→WildCat x) (Cat→WildCat x₁) - + WildCatInstance.wildStrMor CatWS x y = Functor x y + WildCatInstance.toWildFunctor CatWS x y f = + record { F-ob = F-ob ; F-hom = F-hom ; F-id = F-id ; F-seq = F-seq } + where open Functor f + WildCatInstance.mbFunctorApp CatWS = mbFunctorApp + WildCatInstance.F-ty-extractSrc CatWS = + extraxtWildFunSrc >=& λ x → + R.def (quote Cat→WildCat) (R.unknown v∷ R.unknown v∷ v[ x ]) private module C-WS = WildCatInstance CatWS diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index 35ff084f59..968d4a8594 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -54,6 +54,10 @@ FreeTWRM : ∀ {ℓ} (A : Type ℓ) → TypeWithRel ℓ-zero ℓ TypeWithRel.Carrier (FreeTWRM A) = Unit FreeTWRM A TypeWithRel.[ x ∼ x₁ ] = A +extraxtWildFunSrc : R.Term → R.TC (R.Term) +extraxtWildFunSrc (R.con _ l) = matchFirstVarg l +extraxtWildFunSrc (R.def _ l) = matchFirstVarg l +extraxtWildFunSrc t = R.typeError $ "extraxtWildFunSrc fail: " ∷ₑ t ∷ₑ [] @@ -129,16 +133,16 @@ wc→twr : WildCat ℓ ℓ' → TypeWithRel ℓ ℓ' TypeWithRel.Carrier (wc→twr x) = WildCat.ob x TypeWithRel._[_∼_] (wc→twr x) = WildCat.Hom[_,_] x -mbFunctorApp : R.Term → Maybe (R.Term × R.Term) -mbFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t -mbFunctorApp _ = nothing - record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) where no-eta-equality field wildStr : Type (ℓ-suc (ℓ-max ℓ ℓ')) toWildCat : wildStr → WildCat ℓ ℓ' mbIsWildGroupoid : Maybe (∀ WS → IsWildGroupoid (toWildCat WS)) + wildStrMor : wildStr → wildStr → Type (ℓ-max ℓ ℓ') + toWildFunctor : ∀ x y → wildStrMor x y → WildFunctor (toWildCat x) (toWildCat y) + mbFunctorApp : R.Term → Maybe (R.Term × R.Term) + F-ty-extractSrc : R.Term → R.TC R.Term InvFlag = caseMaybe ⊥ Unit mbIsWildGroupoid @@ -170,13 +174,15 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh → (f : ws H[ x , y ]) → FuCases ws {y} {x} (inv ws invF f) ⟅_,_,_⟆FE : ∀ ws {ws' x y} - (F : WildF ws ws') + (F : wildStrMor ws ws') → (f : ws H[ x , y ]) - → FuCases ws' {F-ob F x} {F-ob F y} (F-hom F f ) + → FuCases ws' {F-ob (toWildFunctor _ _ F) x} + {F-ob (toWildFunctor _ _ F) y} + (F-hom (toWildFunctor _ _ F) f ) module _ where open FuExpr' ℓ ℓ' InvFlag wildStr ws→twr - WildF WildFunctor.F-ob public + wildStrMor (λ {x} {y} F → F-ob (toWildFunctor x y F)) public evFuExpr : {ws : wildStr} {x y : ob ws} → FuExpr ws x y → ws H[ x , y ] @@ -184,7 +190,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh evFuExpr {ws} FuExpr'.idFE = id ws evFuExpr {ws} (x FuExpr'.⋆FE x₁) = ws ⟨ evFuExpr x ⋆ evFuExpr x₁ ⟩ evFuExpr {ws} (FuExpr'.invFE invF x) = inv ws invF (evFuExpr x) - evFuExpr FuExpr'.⟅ T , F , x ⟆FE = F ⟪ evFuExpr x ⟫ + evFuExpr FuExpr'.⟅ T , F , x ⟆FE = (toWildFunctor _ _ F) ⟪ evFuExpr x ⟫ module _ {ws : wildStr} where @@ -256,7 +262,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh FuExpr→FF idFE = [ff] FuExpr→FF (x ⋆FE x₁) = (FuExpr→FF x) ff++ (FuExpr→FF x₁) FuExpr→FF (invFE invF x) = ffInv invF (FuExpr→FF x) - FuExpr→FF ⟅ T , F , x ⟆FE = ff⟪ F , FuExpr→FF x ⟫ + FuExpr→FF ⟅ T , F , x ⟆FE = ff⟪ (toWildFunctor _ _ F) , FuExpr→FF x ⟫ evAtom : {ws : wildStr} {x y : ob ws} → FuAtom x y → ws H[ x , y ] @@ -377,8 +383,8 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh (WG ws' invF) F (evFuExpr f) ∙∙ cong (inv ws' invF) (evFF-Func F f) ∙∙ ff⟪⟫inv invF F (FuExpr→FF f) evFF-Func F FuExpr'.⟅ T , F' , f ⟆FE = - cong (F-hom F) (evFF-Func F' f) ∙ - ff-Func F ff⟪ F' , FuExpr→FF f ⟫ + cong (F-hom F) (evFF-Func (toWildFunctor _ _ F') f) ∙ + ff-Func F ff⟪ (toWildFunctor _ _ F') , FuExpr→FF f ⟫ evFuExpr≡evFF : (ws : wildStr) {x y : ob ws} → (f : FuExpr ws x y) → @@ -390,7 +396,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh ∙ evFF++ (FuExpr→FF f) (FuExpr→FF f₁) evFuExpr≡evFF ws (FuExpr'.invFE invF f) = cong (inv ws invF) (evFuExpr≡evFF ws f) ∙ evFFinv invF (FuExpr→FF f) - evFuExpr≡evFF ws FuExpr'.⟅ T , F , f ⟆FE = evFF-Func F f + evFuExpr≡evFF ws FuExpr'.⟅ T , F , f ⟆FE = evFF-Func (toWildFunctor _ _ F) f magicNumber : ℕ magicNumber = 100 @@ -477,6 +483,12 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh λ ws → R.unify (R.def (quote toWildCat) (WS v∷ ws v∷ [])) t >> R.returnTC ws + -- fromWC' : R.Term → R.TC R.Term + -- fromWC' t = tryAllTC + -- (R.typeError ("fromWC fail: " ∷ₑ t ∷ₑ [])) + -- tGs + -- λ ws → R.unify ws t >> R.returnTC ws + tryOp : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) tryOp W zero _ = R.typeError [] tryOp W (suc k) t = do @@ -523,7 +535,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh (R.typeError $ "tryFunc fail " ∷nl t ∷nl t' ∷nl getConTail t') (λ (F-t , x-t) → do F-ty ← R.withNormalisation true $ R.inferType F-t - (W-src , W-trg) ← h F-ty + W-src ← F-ty-extractSrc F-ty R.returnTC {A = R.Term × R.Term × R.Term} (W-src , (F-t , x-t)) ) @@ -536,12 +548,12 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh tm' ← R.checkType tm ty x-t' ← tryE WS-src k x-t R.returnTC (TermExpr.⟅ lift WS-src , F-t , x-t' ⟆FE) - where + -- where - h : R.Term → R.TC (R.Term × R.Term) - h (R.con _ l) = match2Vargs l - h (R.def _ l) = match2Vargs l - h t = R.typeError $ "match2Fail: " ∷ₑ t ∷ₑ [] + -- h : R.Term → R.TC (R.Term × R.Term) + -- h (R.con _ l) = match2Vargs l + -- h (R.def _ l) = match2Vargs l + -- h t = R.typeError $ "match2Fail: " ∷ₑ t ∷ₑ [] @@ -563,6 +575,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh (tryId W t) (R.catchTC (tryInv W k t) (R.catchTC (tryOp W k t) + -- (tryFunc W k t))) (R.catchTC (tryFunc W k t) (atom W t)))) @@ -607,7 +620,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh finalTrm1 = just (R.def (quote evFuExpr≡evFF) (Ws v∷ Wt v∷ expr1 v∷ [])) ∷ invPa1 - + -- R.typeError msg let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms (finalTrm0 ++ symPathTerms finalTrm1) - (R.unify finalTrm hole) + R.catchTC (R.unify finalTrm hole) (R.typeError msg) From 24486fed248711d9f6b69c6deaa1b4b7954217cc Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 15 Apr 2024 06:22:16 +0200 Subject: [PATCH 15/20] passing names of definitioons to ignore --- Cubical/Tactics/GroupSolver/Example.agda | 28 +++++++++++++++++++++++- Cubical/Tactics/GroupSolver/Solver.agda | 17 ++++++++++---- 2 files changed, 40 insertions(+), 5 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index dc9b300b68..900e92144b 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -3,6 +3,7 @@ module Cubical.Tactics.GroupSolver.Example where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms @@ -21,7 +22,7 @@ module example (G G* G○ : Group ℓ) where - open Group-Solver ℓ + open Group-Solver ℓ [] open GroupStr (snd G) @@ -59,3 +60,28 @@ module example (G G* G○ : Group ℓ) lhs≡rhs : lhs ≡ rhs lhs≡rhs = solveGroup (G ∷ G* ∷ G○ ∷ []) + +module ℤexamples where + open import Cubical.Data.Int using (ℤ) + open import Cubical.Algebra.Group.Instances.Int + open import Cubical.Data.Nat using (ℕ) + + + + + open GroupStr (snd ℤGroup) + + module _ k ([_]ᶻ : ℕ → fst ℤGroup) where + + + open Group-Solver ℓ-zero + (quote ℤGroup + ∷ [ quote ℤHom ]) + + + lhs rhs : (fst ℤGroup) + lhs = fst (ℤHom k) ([ 1 ]ᶻ · [ 3 ]ᶻ) + rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ + + lhs≡rhs : lhs ≡ rhs + lhs≡rhs = solveGroup (ℤGroup ∷ []) diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index 01c2ca3121..7d21823e75 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -8,6 +8,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Data.Unit open import Cubical.Data.List +open import Cubical.Data.Bool open import Cubical.Data.Maybe open import Cubical.Data.Sigma @@ -44,7 +45,7 @@ module _ {ℓ} where wildIsIso.sect (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvL f wildIsIso.retr (WildGroupoid.isWildGroupoid Group→WildGroupoid f) = ·InvR f -module Group-Solver ℓ where +module _ ℓ where mbGroupHomApp : R.Term → Maybe (R.Term × R.Term) mbGroupHomApp (R.def (quote fst) t) = match2Vargs' t @@ -75,7 +76,15 @@ module Group-Solver ℓ where private module GRP-WS = WildCatInstance GroupWS - macro - solveGroup : R.Term → R.Term → R.TC Unit - solveGroup = GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) + -- macro + -- solveGroup : R.Term → R.Term → R.TC Unit + -- solveGroup = GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) + + + module Group-Solver (no-norm-defs : List R.Name) where + + macro + solveGroup : R.Term → R.Term → R.TC Unit + solveGroup x y = + R.withReduceDefs (false , no-norm-defs) (GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) x y) From f97cf2b5e7568eb27a56372be01f4530e6109cda Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 15 Apr 2024 10:19:42 +0200 Subject: [PATCH 16/20] only top level WildCat needs to be suplied --- Cubical/Tactics/GroupSolver/Example.agda | 6 ++-- Cubical/Tactics/GroupSolver/Solver.agda | 3 +- Cubical/Tactics/GroupoidSolver/Example.agda | 8 ++--- Cubical/Tactics/GroupoidSolver/Solver.agda | 17 +++++++-- Cubical/Tactics/WildCatSolver/Example.agda | 12 +++---- Cubical/Tactics/WildCatSolver/Solver.agda | 4 +-- Cubical/Tactics/WildCatSolver/Solvers.agda | 40 ++++++++++----------- 7 files changed, 49 insertions(+), 41 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 900e92144b..595179ee23 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -43,7 +43,7 @@ module example (G G* G○ : Group ℓ) pB = ((1g · p) · q) · r pA≡pB : pA ≡ pB - pA≡pB = solveGroup (G ∷ []) + pA≡pB = solveGroup G module T2 p p' q r s t u where @@ -59,7 +59,7 @@ module example (G G* G○ : Group ℓ) lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveGroup (G ∷ G* ∷ G○ ∷ []) + lhs≡rhs = solveGroup G module ℤexamples where open import Cubical.Data.Int using (ℤ) @@ -84,4 +84,4 @@ module ℤexamples where rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveGroup (ℤGroup ∷ []) + lhs≡rhs = solveGroup ℤGroup diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index 7d21823e75..adb5abdf7b 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -70,8 +70,7 @@ module _ ℓ where where open IsGroupHom isGHom WildCatInstance.mbFunctorApp GroupWS = mbGroupHomApp WildCatInstance.F-ty-extractSrc GroupWS = - groupHomTrmSrc >=& λ x → - R.def (quote WildGroupoid.wildCat) v[ R.def (quote Group→WildGroupoid) (v[ x ])] + groupHomTrmSrc private module GRP-WS = WildCatInstance GroupWS diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index 4b07ce4b9d..dbf39e593a 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -51,7 +51,7 @@ module WildGroupoidExamples where pB = p ⋆ (q ⋆ r) pA≡pB : pA ≡ pB - pA≡pB = solveWildGroupoid (WG ∷ WG* ∷ []) + pA≡pB = solveWildGroupoid WG @@ -68,7 +68,7 @@ module WildGroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (WG ∷ WG* ∷ []) + lhs≡rhs = solveWildGroupoid WG module T3 (obs : ℕ → ob) @@ -81,7 +81,7 @@ module WildGroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (WG ∷ []) + lhs≡rhs = solveWildGroupoid WG module GroupoidExamples where @@ -122,4 +122,4 @@ module GroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid (G ∷ G* ∷ []) + lhs≡rhs = solveWildGroupoid G diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index d174eae868..827872fc38 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -29,6 +29,11 @@ module WildGroupoid-Solver ℓ ℓ' where mbWildFunctorApp : R.Term → Maybe (R.Term × R.Term) mbWildFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t mbWildFunctorApp _ = nothing + + unWildCat : R.Term → R.TC R.Term + unWildCat (R.def (quote WildGroupoid.wildCat) t) = matchFirstVarg t + unWildCat _ = R.typeError [ R.strErr "unWildCat fail" ] + GroupoidWS : WildCatInstance ℓ ℓ' @@ -38,7 +43,8 @@ module WildGroupoid-Solver ℓ ℓ' where WildCatInstance.wildStrMor GroupoidWS _ _ = WildFunctor _ _ WildCatInstance.toWildFunctor GroupoidWS _ _ f = f WildCatInstance.mbFunctorApp GroupoidWS = mbWildFunctorApp - WildCatInstance.F-ty-extractSrc GroupoidWS = extraxtWildFunSrc + WildCatInstance.F-ty-extractSrc GroupoidWS = + extraxtWildFunSrc >=> unWildCat private module WGPD-WS = WildCatInstance GroupoidWS @@ -67,6 +73,12 @@ module Groupoid-Solver ℓ ℓ' where wildIsIso.sect wgi = sec wildIsIso.retr wgi = ret + + unCategory : R.Term → R.TC R.Term + unCategory (R.def (quote GroupoidCat.category) t) = matchFirstVarg t + unCategory _ = R.typeError [ R.strErr "unWildCat fail" ] + + GroupoidWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr GroupoidWS = GroupoidCat ℓ ℓ' WildCatInstance.toWildCat GroupoidWS = WildGroupoid.wildCat ∘ Groupoid→WildGroupoid @@ -78,8 +90,7 @@ module Groupoid-Solver ℓ ℓ' where where open Functor f WildCatInstance.mbFunctorApp GroupoidWS = mbFunctorApp WildCatInstance.F-ty-extractSrc GroupoidWS = - extraxtWildFunSrc >=& λ x → - R.def (quote Cat-Solver.Cat→WildCat) (R.unknown v∷ R.unknown v∷ v[ x ]) + extraxtWildFunSrc >=> unCategory private module GPD-WS = WildCatInstance GroupoidWS diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index 22a9c00aeb..8bf9f01df4 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -37,10 +37,10 @@ module exampleWC where open WildCat-Solver ℓ ℓ' pA=pB : pA ≡ pB - pA=pB = solveWildCat (WC ∷ WC* ∷ []) + pA=pB = solveWildCat WC pB=pC : pB ≡ pC - pB=pC = solveWildCat (WC ∷ WC* ∷ []) + pB=pC = solveWildCat WC module exampleC ℓ ℓ' where @@ -70,10 +70,10 @@ module exampleC ℓ ℓ' where pA=pB : pA ≡ pB - pA=pB = solveCat (C ∷ []) + pA=pB = solveCat C pB=pC : pB ≡ pC - pB=pC = solveCat (C ∷ []) + pB=pC = solveCat C module E2 x y z w v (p : Hom[ x , F ⟅ y ⟆ ]) @@ -99,8 +99,8 @@ module exampleC ℓ ℓ' where pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) pA=pB : pA ≡ pB - pA=pB = solveCat (C ∷ C* ∷ []) + pA=pB = solveCat C pB=pC : pB ≡ pC - pB=pC = solveCat (C ∷ C* ∷ []) + pB=pC = solveCat C diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index d56c064e9f..5b5856d196 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -74,9 +74,7 @@ module Cat-Solver ℓ ℓ' where record { F-ob = F-ob ; F-hom = F-hom ; F-id = F-id ; F-seq = F-seq } where open Functor f WildCatInstance.mbFunctorApp CatWS = mbFunctorApp - WildCatInstance.F-ty-extractSrc CatWS = - extraxtWildFunSrc >=& λ x → - R.def (quote Cat→WildCat) (R.unknown v∷ R.unknown v∷ v[ x ]) + WildCatInstance.F-ty-extractSrc CatWS = extraxtWildFunSrc private module C-WS = WildCatInstance CatWS diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index f5364c906d..34f0758b56 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -419,7 +419,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh R.con (quote FuExpr'.⟅_,_,_⟆FE) (T v∷ F v∷ buildFromTE x v∷ []) - module tryWCE WS (tGs : List R.Term) where + module tryWCE WS where mb-invol : R.Term → ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) @@ -476,12 +476,12 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh tryE : (W : R.Term) → ℕ → R.Term → R.TC (TE InvFlag (lift W) _ _) - fromWC : R.Term → R.TC R.Term - fromWC t = tryAllTC - (R.typeError ("fromWC fail: " ∷ₑ t ∷ₑ [])) - tGs - λ ws → R.unify (R.def (quote toWildCat) - (WS v∷ ws v∷ [])) t >> R.returnTC ws + -- fromWC : R.Term → R.TC R.Term + -- fromWC t = tryAllTC + -- (R.typeError ("fromWC fail: " ∷ₑ t ∷ₑ [])) + -- tGs + -- λ ws → R.unify (R.def (quote toWildCat) + -- (WS v∷ ws v∷ [])) t >> R.returnTC ws -- fromWC' : R.Term → R.TC R.Term -- fromWC' t = tryAllTC @@ -531,7 +531,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh tryFunc W (suc k) t = do t' ← R.normalise t -- (R.typeError $ "tryFunc fail " ∷nl t ∷nl t' ∷nl getConTail t') - (WC-src , F-t , x-t) ← Mb.rec + (WS-src , F-t , x-t) ← Mb.rec (R.typeError $ "tryFunc fail " ∷nl t ∷nl t' ∷nl getConTail t') (λ (F-t , x-t) → do F-ty ← R.withNormalisation true $ R.inferType F-t @@ -540,7 +540,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh (W-src , (F-t , x-t)) ) (mbFunctorApp t') - WS-src ← fromWC WC-src + -- WS-src ← {!!} --fromWC WC-src let tm = R.con (quote FuCases.⟅_,_,_⟆FE) (WS-src v∷ F-t v∷ x-t v∷ []) ty = R.def (quote FuCases) @@ -581,20 +581,20 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh solveW : R.Term → R.Term → R.Term → R.TC Unit - solveW Ws Wts' hole = do - Wts ← quotedList→ListOfTerms Wts' - Wt ← tryAllTC - (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) - Wts R.returnTC + solveW Ws Wt hole = do + + -- Wt ← tryAllTC + -- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) + -- Wts R.returnTC hoTy ← R.withNormalisation true $ R.inferType hole >>= wait-for-type (t0 , t1) ← (get-boundary hoTy ) >>= Mb.rec (R.typeError [ R.strErr "unable to get boundary" ]) (λ x → R.returnTC x) - t0' ← tryWCE.tryE Ws Wts Wt magicNumber t0 - t1' ← tryWCE.tryE Ws Wts Wt magicNumber t1 - expr0 ← tryWCE.checkFromTE Ws Wts t0' - expr1 ← tryWCE.checkFromTE Ws Wts t1' + t0' ← tryWCE.tryE Ws Wt magicNumber t0 + t1' ← tryWCE.tryE Ws Wt magicNumber t1 + expr0 ← tryWCE.checkFromTE Ws t0' + expr1 ← tryWCE.checkFromTE Ws t1' let msg = (TermExpr.printFuExpr InvFlag (λ _ → "●") t0' ∷nl TermExpr.printFuExpr InvFlag (λ _ → "●") t1' ∷ₑ []) @@ -602,8 +602,8 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh invol1 ← R.normalise (R.def (quote FuExpr→FF) (Ws v∷ v[ expr1 ])) - red0 ← tryWCE.redList' Ws Wts Wt invol0 - red1 ← tryWCE.redList' Ws Wts Wt invol1 + red0 ← tryWCE.redList' Ws Wt invol0 + red1 ← tryWCE.redList' Ws Wt invol1 let invPa0 = Li.map (λ t' → just (R.def (quote evFF≡↓) (Ws v∷ Wt v∷ t' v∷ []))) From c2052dafbec3ae1e0761aa9841b17778053b5d4f Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 15 Apr 2024 21:16:22 +0200 Subject: [PATCH 17/20] infering top instance --- Cubical/Tactics/GroupSolver/Example.agda | 15 ++++-------- Cubical/Tactics/GroupSolver/Solver.agda | 19 +++++++-------- Cubical/Tactics/GroupoidSolver/Example.agda | 8 +++--- Cubical/Tactics/GroupoidSolver/Solver.agda | 22 ++++++++++++++--- Cubical/Tactics/WildCatSolver/Example.agda | 14 +++++------ Cubical/Tactics/WildCatSolver/Solver.agda | 27 +++++++++++++++++---- Cubical/Tactics/WildCatSolver/Solvers.agda | 18 +++++++++++--- 7 files changed, 79 insertions(+), 44 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 595179ee23..bddaa7c5b2 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -43,7 +43,7 @@ module example (G G* G○ : Group ℓ) pB = ((1g · p) · q) · r pA≡pB : pA ≡ pB - pA≡pB = solveGroup G + pA≡pB = solveGroup module T2 p p' q r s t u where @@ -59,15 +59,11 @@ module example (G G* G○ : Group ℓ) lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveGroup G + lhs≡rhs = solveGroup module ℤexamples where - open import Cubical.Data.Int using (ℤ) - open import Cubical.Algebra.Group.Instances.Int open import Cubical.Data.Nat using (ℕ) - - - + open import Cubical.Algebra.Group.Instances.Int open GroupStr (snd ℤGroup) @@ -75,8 +71,7 @@ module ℤexamples where open Group-Solver ℓ-zero - (quote ℤGroup - ∷ [ quote ℤHom ]) + (quote ℤGroup ∷ [ quote ℤHom ]) lhs rhs : (fst ℤGroup) @@ -84,4 +79,4 @@ module ℤexamples where rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveGroup ℤGroup + lhs≡rhs = solveGroup[ ℤGroup ] diff --git a/Cubical/Tactics/GroupSolver/Solver.agda b/Cubical/Tactics/GroupSolver/Solver.agda index adb5abdf7b..c08b1dec92 100644 --- a/Cubical/Tactics/GroupSolver/Solver.agda +++ b/Cubical/Tactics/GroupSolver/Solver.agda @@ -69,21 +69,20 @@ module _ ℓ where record { F-ob = _ ; F-hom = f ; F-id = pres1 ; F-seq = pres· } where open IsGroupHom isGHom WildCatInstance.mbFunctorApp GroupWS = mbGroupHomApp - WildCatInstance.F-ty-extractSrc GroupWS = - groupHomTrmSrc - + WildCatInstance.F-ty-extractSrc GroupWS = groupHomTrmSrc + WildCatInstance.extractWS GroupWS = unFst private module GRP-WS = WildCatInstance GroupWS - -- macro - -- solveGroup : R.Term → R.Term → R.TC Unit - -- solveGroup = GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) - module Group-Solver (no-norm-defs : List R.Name) where macro - solveGroup : R.Term → R.Term → R.TC Unit - solveGroup x y = - R.withReduceDefs (false , no-norm-defs) (GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) x y) + solveGroup[_] : R.Term → R.Term → R.TC Unit + solveGroup[ x ] y = + R.withReduceDefs (false , no-norm-defs) (GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) (just x) y) + + solveGroup : R.Term → R.TC Unit + solveGroup y = + R.withReduceDefs (false , no-norm-defs) (GRP-WS.solveW (R.def (quote GroupWS) ( R.unknown v∷ [])) nothing y) diff --git a/Cubical/Tactics/GroupoidSolver/Example.agda b/Cubical/Tactics/GroupoidSolver/Example.agda index dbf39e593a..a34af83fe0 100644 --- a/Cubical/Tactics/GroupoidSolver/Example.agda +++ b/Cubical/Tactics/GroupoidSolver/Example.agda @@ -51,7 +51,7 @@ module WildGroupoidExamples where pB = p ⋆ (q ⋆ r) pA≡pB : pA ≡ pB - pA≡pB = solveWildGroupoid WG + pA≡pB = solveWildGroupoid @@ -68,7 +68,7 @@ module WildGroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid WG + lhs≡rhs = solveWildGroupoid module T3 (obs : ℕ → ob) @@ -81,7 +81,7 @@ module WildGroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid WG + lhs≡rhs = solveWildGroupoid module GroupoidExamples where @@ -122,4 +122,4 @@ module GroupoidExamples where lhs≡rhs : lhs ≡ rhs - lhs≡rhs = solveWildGroupoid G + lhs≡rhs = solveWildGroupoid diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index 827872fc38..effd307396 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -45,12 +45,20 @@ module WildGroupoid-Solver ℓ ℓ' where WildCatInstance.mbFunctorApp GroupoidWS = mbWildFunctorApp WildCatInstance.F-ty-extractSrc GroupoidWS = extraxtWildFunSrc >=> unWildCat + WildCatInstance.extractWS GroupoidWS = + WildCat-Solver.extrWS ℓ ℓ' >=> unWildCat + private module WGPD-WS = WildCatInstance GroupoidWS macro - solveWildGroupoid : R.Term → R.Term → R.TC Unit - solveWildGroupoid = WGPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) + + solveWildGroupoid : R.Term → R.TC Unit + solveWildGroupoid = WGPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) nothing + + solveWildGroupoid[_] : R.Term → R.Term → R.TC Unit + solveWildGroupoid[ x ] = WGPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) (just x) + module Groupoid-Solver ℓ ℓ' where open import Cubical.Categories.Functor @@ -91,10 +99,16 @@ module Groupoid-Solver ℓ ℓ' where WildCatInstance.mbFunctorApp GroupoidWS = mbFunctorApp WildCatInstance.F-ty-extractSrc GroupoidWS = extraxtWildFunSrc >=> unCategory + WildCatInstance.extractWS GroupoidWS = + Cat-Solver.extrWS ℓ ℓ' >=> unCategory private module GPD-WS = WildCatInstance GroupoidWS macro - solveWildGroupoid : R.Term → R.Term → R.TC Unit - solveWildGroupoid = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) + solveWildGroupoid : R.Term → R.TC Unit + solveWildGroupoid = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) nothing + + solveWildGroupoid[_] : R.Term → R.Term → R.TC Unit + solveWildGroupoid[ x ] = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) (just x) + diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index 8bf9f01df4..7be6d1cd75 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -37,10 +37,10 @@ module exampleWC where open WildCat-Solver ℓ ℓ' pA=pB : pA ≡ pB - pA=pB = solveWildCat WC - + pA=pB = solveWildCat + pB=pC : pB ≡ pC - pB=pC = solveWildCat WC + pB=pC = solveWildCat module exampleC ℓ ℓ' where @@ -70,10 +70,10 @@ module exampleC ℓ ℓ' where pA=pB : pA ≡ pB - pA=pB = solveCat C + pA=pB = solveCat pB=pC : pB ≡ pC - pB=pC = solveCat C + pB=pC = solveCat module E2 x y z w v (p : Hom[ x , F ⟅ y ⟆ ]) @@ -99,8 +99,8 @@ module exampleC ℓ ℓ' where pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) pA=pB : pA ≡ pB - pA=pB = solveCat C + pA=pB = solveCat pB=pC : pB ≡ pC - pB=pC = solveCat C + pB=pC = solveCat diff --git a/Cubical/Tactics/WildCatSolver/Solver.agda b/Cubical/Tactics/WildCatSolver/Solver.agda index 5b5856d196..18b80c67c6 100644 --- a/Cubical/Tactics/WildCatSolver/Solver.agda +++ b/Cubical/Tactics/WildCatSolver/Solver.agda @@ -28,6 +28,9 @@ module WildCat-Solver ℓ ℓ' where mbWildFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t mbWildFunctorApp _ = nothing + extrWS : R.Term → R.TC R.Term + extrWS (R.def (quote WildCat.Hom[_,_]) t) = matchFirstVarg t + extrWS t = R.typeError (R.strErr "extrWS fail : " ∷ [ R.termErr t ]) WildCatWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr (WildCatWS) = WildCat ℓ ℓ' @@ -37,13 +40,17 @@ module WildCat-Solver ℓ ℓ' where WildCatInstance.toWildFunctor WildCatWS _ _ f = f WildCatInstance.mbFunctorApp WildCatWS = mbWildFunctorApp WildCatInstance.F-ty-extractSrc WildCatWS = extraxtWildFunSrc + WildCatInstance.extractWS WildCatWS = extrWS private module WC-WS = WildCatInstance WildCatWS macro - solveWildCat : R.Term → R.Term → R.TC Unit - solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) + solveWildCat[_] : R.Term → R.Term → R.TC Unit + solveWildCat[_] x = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) (just x) + + solveWildCat : R.Term → R.TC Unit + solveWildCat = WC-WS.solveW (R.def (quote WildCatWS) ( R.unknown v∷ R.unknown v∷ [])) nothing module Cat-Solver ℓ ℓ' where @@ -65,6 +72,11 @@ module Cat-Solver ℓ ℓ' where WildCat.⋆Assoc (Cat→WildCat x) = Category.⋆Assoc x + extrWS : R.Term → R.TC R.Term + extrWS (R.def (quote Category.Hom[_,_]) t) = matchFirstVarg t + extrWS _ = R.typeError [ R.strErr "extrWS fail" ] + + CatWS : WildCatInstance ℓ ℓ' WildCatInstance.wildStr CatWS = Category ℓ ℓ' WildCatInstance.toWildCat CatWS = Cat→WildCat @@ -74,10 +86,15 @@ module Cat-Solver ℓ ℓ' where record { F-ob = F-ob ; F-hom = F-hom ; F-id = F-id ; F-seq = F-seq } where open Functor f WildCatInstance.mbFunctorApp CatWS = mbFunctorApp - WildCatInstance.F-ty-extractSrc CatWS = extraxtWildFunSrc + WildCatInstance.F-ty-extractSrc CatWS = extraxtWildFunSrc + WildCatInstance.extractWS CatWS = extrWS + private module C-WS = WildCatInstance CatWS macro - solveCat : R.Term → R.Term → R.TC Unit - solveCat = C-WS.solveW (R.def (quote CatWS) ( R.unknown v∷ R.unknown v∷ [])) + solveCat[_] : R.Term → R.Term → R.TC Unit + solveCat[_] x = C-WS.solveW (R.def (quote CatWS) ( R.unknown v∷ R.unknown v∷ [])) (just x) + + solveCat : R.Term → R.TC Unit + solveCat = C-WS.solveW (R.def (quote CatWS) ( R.unknown v∷ R.unknown v∷ [])) nothing diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index 34f0758b56..82f59b77b3 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -143,6 +143,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh toWildFunctor : ∀ x y → wildStrMor x y → WildFunctor (toWildCat x) (toWildCat y) mbFunctorApp : R.Term → Maybe (R.Term × R.Term) F-ty-extractSrc : R.Term → R.TC R.Term + extractWS : R.Term → R.TC R.Term InvFlag = caseMaybe ⊥ Unit mbIsWildGroupoid @@ -580,17 +581,26 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh - solveW : R.Term → R.Term → R.Term → R.TC Unit - solveW Ws Wt hole = do + solveW : R.Term → Maybe R.Term → R.Term → R.TC Unit + solveW Ws mbWt hole = do -- Wt ← tryAllTC -- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) -- Wts R.returnTC + + hoTy ← R.withNormalisation true $ R.inferType hole >>= wait-for-type - (t0 , t1) ← (get-boundary hoTy ) >>= Mb.rec + (A , (t0 , t1)) ← (get-boundaryWithDom hoTy ) >>= Mb.rec (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) + pure + + A' ← R.normalise A >>= wait-for-type + + Wt ← Mb.rec + (extractWS A') + pure mbWt + t0' ← tryWCE.tryE Ws Wt magicNumber t0 t1' ← tryWCE.tryE Ws Wt magicNumber t1 expr0 ← tryWCE.checkFromTE Ws t0' From 23b1ddf10645ee003fcd9a34c7afc1161ef88f99 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 15 Apr 2024 21:16:49 +0200 Subject: [PATCH 18/20] whitespace fix --- Cubical/Tactics/GroupSolver/Example.agda | 4 ++-- Cubical/Tactics/GroupoidSolver/Solver.agda | 5 ++--- Cubical/Tactics/WildCatSolver/Example.agda | 2 +- Cubical/Tactics/WildCatSolver/Solvers.agda | 4 ++-- 4 files changed, 7 insertions(+), 8 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index bddaa7c5b2..66f6724c30 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -75,8 +75,8 @@ module ℤexamples where lhs rhs : (fst ℤGroup) - lhs = fst (ℤHom k) ([ 1 ]ᶻ · [ 3 ]ᶻ) - rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ + lhs = fst (ℤHom k) ([ 1 ]ᶻ · [ 3 ]ᶻ) + rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ lhs≡rhs : lhs ≡ rhs lhs≡rhs = solveGroup[ ℤGroup ] diff --git a/Cubical/Tactics/GroupoidSolver/Solver.agda b/Cubical/Tactics/GroupoidSolver/Solver.agda index effd307396..473b40cb0e 100644 --- a/Cubical/Tactics/GroupoidSolver/Solver.agda +++ b/Cubical/Tactics/GroupoidSolver/Solver.agda @@ -29,7 +29,7 @@ module WildGroupoid-Solver ℓ ℓ' where mbWildFunctorApp : R.Term → Maybe (R.Term × R.Term) mbWildFunctorApp (R.def (quote WildFunctor.F-hom) t) = matchFunctorAppArgs t mbWildFunctorApp _ = nothing - + unWildCat : R.Term → R.TC R.Term unWildCat (R.def (quote WildGroupoid.wildCat) t) = matchFirstVarg t unWildCat _ = R.typeError [ R.strErr "unWildCat fail" ] @@ -47,7 +47,7 @@ module WildGroupoid-Solver ℓ ℓ' where extraxtWildFunSrc >=> unWildCat WildCatInstance.extractWS GroupoidWS = WildCat-Solver.extrWS ℓ ℓ' >=> unWildCat - + private module WGPD-WS = WildCatInstance GroupoidWS @@ -111,4 +111,3 @@ module Groupoid-Solver ℓ ℓ' where solveWildGroupoid[_] : R.Term → R.Term → R.TC Unit solveWildGroupoid[ x ] = GPD-WS.solveW (R.def (quote GroupoidWS) ( R.unknown v∷ R.unknown v∷ [])) (just x) - diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index 7be6d1cd75..ef33a19ef0 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -38,7 +38,7 @@ module exampleWC where pA=pB : pA ≡ pB pA=pB = solveWildCat - + pB=pC : pB ≡ pC pB=pC = solveWildCat diff --git a/Cubical/Tactics/WildCatSolver/Solvers.agda b/Cubical/Tactics/WildCatSolver/Solvers.agda index 82f59b77b3..e26cbda773 100644 --- a/Cubical/Tactics/WildCatSolver/Solvers.agda +++ b/Cubical/Tactics/WildCatSolver/Solvers.agda @@ -583,7 +583,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh solveW : R.Term → Maybe R.Term → R.Term → R.TC Unit solveW Ws mbWt hole = do - + -- Wt ← tryAllTC -- (R.typeError $ "At least one 𝑿 must be provded!" ∷ₑ []) -- Wts R.returnTC @@ -595,7 +595,7 @@ record WildCatInstance ℓ ℓ' : Type (ℓ-suc (ℓ-suc (ℓ-max ℓ ℓ'))) wh (R.typeError [ R.strErr "unable to get boundary" ]) pure - A' ← R.normalise A >>= wait-for-type + A' ← R.normalise A >>= wait-for-type Wt ← Mb.rec (extractWS A') From c880847a66735c80d0cdb4bf977db77d02024550 Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 2 Sep 2024 23:01:35 +0200 Subject: [PATCH 19/20] removed obsolete Path Solver --- Cubical/Tactics/PathSolver/Example.agda | 85 --- Cubical/Tactics/PathSolver/Solver.agda | 796 ------------------------ 2 files changed, 881 deletions(-) delete mode 100644 Cubical/Tactics/PathSolver/Example.agda delete mode 100644 Cubical/Tactics/PathSolver/Solver.agda diff --git a/Cubical/Tactics/PathSolver/Example.agda b/Cubical/Tactics/PathSolver/Example.agda deleted file mode 100644 index 8b8862dd2f..0000000000 --- a/Cubical/Tactics/PathSolver/Example.agda +++ /dev/null @@ -1,85 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Example where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Function - - -open import Cubical.Tactics.PathSolver.Solver - -private - variable - ℓ : Level - A B C : Type ℓ - -module Examples (x y z w : A) (p p' : x ≡ y) (q : y ≡ z) (q' : z ≡ y) (r : w ≡ z) where - - pA pB pC : x ≡ w - pA = (p ∙∙ refl ∙∙ q) ∙' sym r - pB = (p ∙ (q ∙ sym (sym r ∙ r) ∙ sym q) ∙∙ q ∙∙ refl) ∙∙ sym r ∙∙ refl - pC = refl ∙∙ p' ∙ q ∙ sym q ∙ sym p' ∙∙ p ∙∙ q ∙∙ sym q ∙ q ∙ sym r - - pA=pB : pA ≡ pB - pA=pB = ≡! - - pB=pC : pB ≡ pC - pB=pC = ≡! - - pA=pC : pA ≡ pC - pA=pC = ≡! - - sqTest : Square p (sym r ∙ refl) (p ∙ q) (q ∙ sym r) - sqTest = sq! - - -module Examples2 (x y z : B) (w : A) (f g : B → A) - (p p' : x ≡ y) (q : y ≡ z) (r : w ≡ f z) where - - pA pB : f x ≡ w - pA = cong f (p' ∙ sym p') ∙' cong f (p ∙∙ refl ∙∙ q) ∙ sym r - pB = (cong f p ∙ (cong f q ∙ (sym (sym r ∙ (refl ∙ refl) ∙ refl ∙ r)) ∙ cong f (sym q)) ∙∙ cong f q ∙∙ refl) ∙∙ sym r ∙∙ refl - - - pA=pB : pA ≡ pB - pA=pB = ≡!cong f - - - -module Examples3 (x y z : B) (w : A) (f : B → A) - (p p' : x ≡ y) (q q' : y ≡ z) (r r' : f z ≡ w) where - - squares≃Example : (Square (cong f p') (cong f q ∙ r) (cong f p) (cong f q' ∙ r')) - ≃ (cong f (sym (p ∙ q) ∙∙ p' ∙∙ q') ≡ r ∙ sym r') - - squares≃Example = - 2-cylinder-from-square.Sq≃Sq' - (cong f ((p ∙ q))) - (≡!cong f) - - -module Examples4 (f : B → A) (g : C → B) - (x y : B) (z w v : C) - (p : x ≡ y) (q : y ≡ g z) (r : z ≡ w) (u : w ≡ v) where - - - sqE : (cong f p ∙∙ cong f q ∙∙ cong (f ∘ g) r) ∙ (cong (f ∘ g) u) - ≡ (cong f (p ∙ q)) ∙ (cong f (cong g r ∙ cong g u )) - sqE = ≡!cong (f ∷ g ∷ [fn]) - - - sqE' : Square (cong f p ∙∙ cong f q ∙∙ cong (f ∘ g) r) - (cong f (cong g r ∙ cong g u )) - (cong f (p ∙ q)) - (cong (f ∘ g) u) - sqE' = sq!cong (f ∷ g ∷ [fn]) - - -module _ {A : Type ℓ} {x y z : A} (p q : x ≡ x) where - - SqCompEquiv : (Square p p q q) ≃ (p ∙ q ≡ q ∙ p) - SqCompEquiv = 2-cylinder-from-square.Sq≃Sq' refl ≡! diff --git a/Cubical/Tactics/PathSolver/Solver.agda b/Cubical/Tactics/PathSolver/Solver.agda deleted file mode 100644 index b040b784fe..0000000000 --- a/Cubical/Tactics/PathSolver/Solver.agda +++ /dev/null @@ -1,796 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Tactics.PathSolver.Solver where - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path - -open import Cubical.Data.Bool as 𝟚 hiding (_≤_) -open import Cubical.Data.Nat as ℕ hiding (_·_) -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.List as Li -open import Cubical.Data.Maybe as Mb - - -open import Cubical.HITs.Interval - -open import Cubical.Relation.Nullary - -open import Cubical.Reflection.Base -import Agda.Builtin.Reflection as R -open import Cubical.Tactics.WildCatSolver.Reflection -open import Cubical.Tactics.Reflection -open import Agda.Builtin.String - - -private - variable - ℓ ℓ' : Level - A A' : Type ℓ - - -infixr 5 _∷Tω_ - -infixr 5 _∷_ - - -typeView : R.Term → R.TC (Maybe R.Term) -typeView t = - ⦇ just (R.checkType t (R.def (quote Type) v[ R.unknown ])) ⦈ <|> ⦇ nothing ⦈ - - - --- macro --- tvTest : R.Term → R.Term → R.TC Unit --- tvTest ty? hole = --- typeView ty? >>= --- Mb.caseMaybe --- (R.typeError [ "not a type!" ]ₑ) --- (R.typeError [ "is type!" ]ₑ) - - -data [Typeₙ] : Typeω where - [Tω] : [Typeₙ] - _∷Tω_ : ∀ {ℓ} → Type ℓ → [Typeₙ] → [Typeₙ] - -data [Fns] : Typeω where - [fn] : [Fns] - _∷_ : ∀ {ℓ ℓ'} → {A : Type ℓ} {A' : Type ℓ'} → (A → A') → [Fns] → [Fns] - - -gatherCodoms : R.Term → R.TC (List R.Term) -gatherCodoms t@(R.con (quote [fn]) _) = pure [] -gatherCodoms t@(R.con (quote [Fns]._∷_) (_ h∷ _ h∷ a' h∷ _ h∷ _ v∷ xs v∷ [])) = - ⦇ ⦇ a' ⦈ ∷ gatherCodoms xs ⦈ -gatherCodoms t = R.typeError $ "gatherCodoms - FAIL : " ∷ₑ [ t ]ₑ - -inferTypesFromFns : R.Term → R.TC (List R.Term) -inferTypesFromFns = - (wait-for-term >=> gatherCodoms) >=> uniqeAtoms - -reflected[Typeₙ]→[reflectedTy] : R.Term → R.TC (List R.Term) -reflected[Typeₙ]→[reflectedTy] (R.con (quote [Tω]) args) = pure [] -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] (R.con (quote _∷Tω_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Typeₙ]→[reflectedTy] x₁ -reflected[Typeₙ]→[reflectedTy] _ = R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ - -- typeView x >>= - -- Mb.rec - -- (R.typeError [ "reflected[Typeₙ]→[reflectedTy] - FAIL" ]ₑ) - -- (pure ∘ [_]) - -[reflectedTy]→reflected[Typeₙ] : List R.Term → R.Term -[reflectedTy]→reflected[Typeₙ] [] = (R.con (quote [Tω]) []) -[reflectedTy]→reflected[Typeₙ] (x ∷ xs) = - (R.con (quote _∷Tω_) (x v∷ [reflectedTy]→reflected[Typeₙ] xs v∷ [])) - -castTo[Typeₙ] : R.Term → R.TC (R.Term) -castTo[Typeₙ] t@(R.con (quote [Tω]) _) = pure t -castTo[Typeₙ] t@(R.con (quote _∷Tω_) _) = pure t -castTo[Typeₙ] x = - typeView x >>= - Mb.rec - (R.typeError [ "castTo[Typeₙ] - FAIL" ]ₑ) - (λ t → pure (R.con (quote _∷Tω_) (t v∷ (R.con (quote [Tω]) []) v∷ []))) - -castTo[Fns] : R.Term → R.TC (R.Term) -castTo[Fns] t@(R.con (quote [fn]) _) = pure t -castTo[Fns] t@(R.con (quote [Fns]._∷_) _) = pure t -castTo[Fns] x = - R.checkType - (R.con (quote [Fns]._∷_) (x v∷ (R.con (quote [fn]) []) v∷ [])) - (R.def (quote [Fns]) []) - -- funView x >>= - -- Mb.rec - -- (R.typeError [ "castTo[Fns] - FAIL" ]ₑ) - -- (λ t → pure (R.con (quote [Fns]._∷_) (t v∷ (R.con (quote [fn]) []) v∷ []))) - - - --- macro --- test-refl[T]macro : R.Term → R.Term → R.TC Unit --- test-refl[T]macro inp hole = do --- l ← reflected[Typeₙ]→[reflectedTy] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[T] : Unit --- test-refl[T] = test-refl[T]macro (Type ℓ ∷Tω (Type ℓ' → Type) ∷Tω Type₂ ∷Tω []) - - -reflected[Fns]→[reflectedFn] : R.Term → R.TC (List R.Term) -reflected[Fns]→[reflectedFn] (R.con (quote [fn]) args) = pure [] -reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ _ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (_ h∷ x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] (R.con (quote [Fns]._∷_) (x v∷ x₁ v∷ [])) = - (x ∷_) <$> reflected[Fns]→[reflectedFn] x₁ -reflected[Fns]→[reflectedFn] _ = R.typeError [ "reflected[Fns]→[reflectedFn] - FAIL" ]ₑ - --- macro --- test-refl[Fn]macro : R.Term → R.Term → R.TC Unit --- test-refl[Fn]macro inp hole = do --- l ← reflected[Fns]→[reflectedFn] inp --- R.typeError (niceAtomList l) - --- module _ (ℓ ℓ' : Level) where --- test-refl[Fn] : Unit --- test-refl[Fn] = test-refl[Fn]macro --- (suc ∷fn (λ (T : Type ℓ') → (T → Type ℓ)) ∷fn [fn]) - - -maxℓ : [Typeₙ] → Level -maxℓ [Tω] = ℓ-zero -maxℓ (_∷Tω_ {ℓ} _ x₁) = ℓ-max ℓ (maxℓ x₁) - -lookupTₙ : (Ts : [Typeₙ]) → ℕ → Type (maxℓ Ts) -lookupTₙ [Tω] x = ⊥* -lookupTₙ (x₁ ∷Tω Ts) zero = Lift {_} {maxℓ Ts} x₁ -lookupTₙ (_∷Tω_ {ℓ} x₁ Ts) (suc x) = Lift {_} {ℓ} (lookupTₙ Ts x) - -ℓAt : (Ts : [Typeₙ]) → ℕ → Level -ℓAt [Tω] x = ℓ-zero -ℓAt (_∷Tω_ {ℓ} x₁ Ts) zero = ℓ -ℓAt (x₁ ∷Tω Ts) (suc x) = ℓAt Ts x - -TyAt : (Ts : [Typeₙ]) → ∀ k → Type (ℓAt Ts k) -TyAt [Tω] k = ⊥* -TyAt (x ∷Tω Ts) zero = x -TyAt (x ∷Tω Ts) (suc k) = TyAt Ts k - -cast↓ : ∀ Ts k → lookupTₙ Ts k → TyAt Ts k -cast↓ [Tω] k () -cast↓ (x₁ ∷Tω Ts) zero x = lower x -cast↓ (x₁ ∷Tω Ts) (suc k) x = cast↓ Ts k (lower x) - -cast↓Inj : ∀ {[T] A x y} → cast↓ [T] A x ≡ cast↓ [T] A y → x ≡ y -cast↓Inj {[Tω]} {A = A} {()} -cast↓Inj {_ ∷Tω [T]} {A = zero} {lift lower₁} {lift lower₂} = cong lift -cast↓Inj {_ ∷Tω [T]} {A = suc A} {lift lower₁} {lift lower₂} p = - cong lift (cast↓Inj {[T] = [T]} {A = A} p) - -cast↓Inj-sec : ∀ {[T] A x y} p → - cast↓Inj {[T]} {A} {x} {y} (cong (cast↓ [T] A ) p) ≡ p -cast↓Inj-sec {x ∷Tω [T]} {A = zero} p = refl -cast↓Inj-sec {x ∷Tω [T]} {A = suc A} p = - cong (cong lift) $ cast↓Inj-sec {[T]} {A} (cong lower p) - -cast↓Inj-∙∙ : ∀ {[T] A x y z w} p q r → - cast↓Inj {[T]} {A} {x} {w} (p ∙∙ q ∙∙ r) - ≡ (cast↓Inj p ∙∙ cast↓Inj {[T]} {A} {y} {z} q ∙∙ cast↓Inj r) - - -cast↓Inj-∙∙ {x ∷Tω [T]} {zero} p q r = cong-∙∙ lift _ _ _ -cast↓Inj-∙∙ {_ ∷Tω [T]} {suc A} p q r = - cong (cong lift) (cast↓Inj-∙∙ {[T]} {A} p q r) ∙ cong-∙∙ lift _ _ _ - - -cast↑ : ∀ Ts k → TyAt Ts k → lookupTₙ Ts k -cast↑ [Tω] k () -cast↑ (x₁ ∷Tω Ts) zero x = lift x -cast↑ (x₁ ∷Tω Ts) (suc k) x = lift $ cast↑ Ts k x - --- Ts-arr : (Ts : [Typeₙ]) → ∀ s t → Type (ℓ-max (ℓAt Ts s) (ℓAt Ts t)) --- Ts-arr Ts s t = TyAt Ts s → TyAt Ts t - - --- Ts-arr' : (Ts : [Typeₙ]) → ℕ → ℕ → Type (maxℓ Ts) --- Ts-arr' [] x x₁ = Unit --- Ts-arr' (x₂ ∷Tω Ts) zero zero = Lift {_} {maxℓ Ts} (x₂ → x₂) --- Ts-arr' (x₂ ∷Tω Ts) zero (suc x₁) = {!Ts!} --- Ts-arr' (x₂ ∷Tω Ts) (suc x) zero = {!!} --- Ts-arr' (_∷Tω_ {ℓ} x₂ Ts) (suc x) (suc x₁) = --- Lift {_} {ℓ} (Ts-arr' (Ts) (x) (x₁)) - --- Ts-arr' : (Ts : [Typeₙ]) → ∀ s t → --- (lookupTₙ Ts s → lookupTₙ Ts t) → Ts-arr Ts s t --- Ts-arr' Ts s t x x₁ = {!!} - - - -data PathCase : {ℓ : Level} {A : Type ℓ} {a₀ a₁ : A} → a₀ ≡ a₁ → Typeω where - reflCase : ∀ {ℓ A x} → PathCase {ℓ} {A} (refl {x = x}) - compCase : ∀ {ℓ A x y z w} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → PathCase {ℓ} {A = A} (p ∙∙ q ∙∙ r) - congCase : ∀ {ℓ ℓ' A A'} {x y} (f : A → A') - → (p : Path {ℓ} A x y) → PathCase {ℓ'} {A = A'} (cong f p) - - -module _ {ℓ ℓ'} {A : Type ℓ} {A' : Type ℓ'} (f : A → A') {x y} - (p : Path {ℓ} A x y) where - - -- pathCaseCongTest : PathCase λ i → f (p i) - -- pathCaseCongTest = congCase f {!!} - - -infixl 5 _fp∷_ -infixl 5 _fp++_ - -data FlatPath {ℓ} (A : Type ℓ) : Bool → (a₀ a₁ : A) → Type ℓ where - [fp] : ∀ {x b} → FlatPath A b x x - _fp∷_ : ∀ {x y z b} → FlatPath A b x y → y ≡ z → FlatPath A b x z - _invol∷_ : ∀ {x y z} → FlatPath A true x y → y ≡ z → FlatPath A true x y - -magicNumber = 100 - -mb-invol : ℕ → R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) -mb-invol _ (R.con (quote [fp]) _) = R.returnTC nothing -mb-invol (suc n) (R.con (quote _fp∷_) tl) = match2Vargs tl >>= w - where - w : (R.Term × R.Term) → R.TC (Maybe (R.Term × R.Term)) - w (R.con (quote [fp]) _ , _) = R.returnTC nothing - w (xs'@(R.con (quote _fp∷_) tl') , y) = - match2Vargs tl' >>= λ (xs , x) → - R.catchTC - (R.noConstraints $ R.unify - (R.def (quote sym) (x v∷ [])) y - >> (Mb.rec (xs , xs) (idfun _) <$> mb-invol n xs) - >>= λ (xs* , xs*↑) → - R.returnTC - (just ((R.con (quote _invol∷_) (xs* v∷ x v∷ [])) - , xs*↑ ) - )) - (map-Maybe (map-both (λ xs* → R.con (quote _fp∷_) - ((xs* v∷ y v∷ [])))) - <$> mb-invol n xs') - w (x , y) = R.typeError ("Imposible!!₁ : " ∷ₑ x ∷ₑ "\n\n " ∷ₑ y ∷ₑ []) -mb-invol _ x = R.typeError ("Imposible!!₂ : " ∷ₑ x ∷ₑ []) - -mb-invol' : R.Term → R.TC (Maybe (R.Term × R.Term)) -mb-invol' = mb-invol magicNumber - - -redList : ℕ → R.Term → R.TC (List R.Term) -redList = h - where - h : ℕ → R.Term → R.TC (List R.Term) - h zero _ = R.typeError ("magic number too low in mb-invol" ∷ₑ []) - h (suc k) t = - (mb-invol' t) >>= - Mb.rec - (R.returnTC []) - λ (t' , t'↓) → do - p' ← h k t'↓ - R.returnTC (t' ∷ p') - - -redList' : R.Term → R.TC (List R.Term) -redList' = redList magicNumber - - -pattern fp[_] x = [fp] fp∷ x - -FP⟦_⟧ : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → a₀ ≡ a₁ -FP⟦ [fp] ⟧ = refl -FP⟦ x fp∷ x₁ ⟧ = FP⟦ x ⟧ ∙ x₁ - --- FP⟦_⟧t : {a₀ a₁ : A} → FlatPath A true a₀ a₁ → a₀ ≡ a₁ --- FP⟦ [] ⟧t = refl --- FP⟦ x fp∷ x₁ ⟧t = FP⟦ x ⟧t ∙ x₁ --- FP⟦ x invol∷ x₁ ⟧t = (FP⟦ x ⟧t ∙ x₁) ∙ sym x₁ - - -fpF→T : {a₀ a₁ : A} → FlatPath A false a₀ a₁ → FlatPath A true a₀ a₁ -fpF→T [fp] = [fp] -fpF→T (x fp∷ x₁) = fpF→T x fp∷ x₁ - -fpT→F : {a₀ a₁ : A} → Bool → FlatPath A true a₀ a₁ → FlatPath A false a₀ a₁ -fpT→F _ [fp] = [fp] -fpT→F b (x₁ fp∷ x₂) = fpT→F b x₁ fp∷ x₂ -fpT→F false (x₁ invol∷ x₂) = fpT→F false x₁ fp∷ x₂ fp∷ sym x₂ -fpT→F true (x₁ invol∷ x₂) = fpT→F true x₁ - -fpT≡F : {a₀ a₁ : A} → (fp : FlatPath A true a₀ a₁) → - FP⟦ fpT→F false fp ⟧ ≡ FP⟦ fpT→F true fp ⟧ -fpT≡F [fp] = refl -fpT≡F (fp fp∷ x) i = fpT≡F fp i ∙ x -fpT≡F {a₀ = a₀} {a₁} (fp invol∷ x) i j = - hcomp - (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ k ∧ ~ i) - }) - (hcomp (λ k → λ { (i = i1) → FP⟦ fpT→F true fp ⟧ j - ; (j = i0) → a₀ - ; (j = i1) → x (~ i ∧ k) - }) - (fpT≡F fp i j)) - -_fp++_ : ∀ {x y z} - → FlatPath A false x y - → FlatPath A false y z - → FlatPath A false x z -x fp++ [fp] = x -x fp++ (x₁ fp∷ x₂) = x fp++ x₁ fp∷ x₂ - -fp++∙ : {a₀ a₁ a₂ : A} → (fp : FlatPath A false a₀ a₁) - (fp' : FlatPath A false a₁ a₂) - → FP⟦ fp ⟧ ∙ FP⟦ fp' ⟧ ≡ FP⟦ fp fp++ fp' ⟧ -fp++∙ fp [fp] = sym (rUnit _) -fp++∙ fp (fp' fp∷ x) = assoc _ _ _ ∙ cong (_∙ x) (fp++∙ fp fp') - -module PE ([T] : [Typeₙ]) where - - data PathExpr : (k : ℕ) (a₀ a₁ : lookupTₙ [T] k) → Type (maxℓ [T]) where - reflExpr : ∀ {A x} → PathExpr A x x - atomExpr : ∀ {A x y} → (p : x ≡ y) → PathExpr A x y - compExpr : ∀ {A x y z w} - → PathExpr A x y → PathExpr A y z → PathExpr A z w - → PathExpr A x w - congExpr : ∀ {A A' x y} → (f : lookupTₙ [T] A → lookupTₙ [T] A') - → PathExpr A x y - → PathExpr A' (f x) (f y) - - - PE⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ → - (cast↓ [T] A a₀) ≡ (cast↓ [T] A a₁) - PE⟦ reflExpr ⟧ = refl - PE⟦ atomExpr p ⟧ = cong _ p - PE⟦ compExpr x x₁ x₂ ⟧ = PE⟦ x ⟧ ∙∙ PE⟦ x₁ ⟧ ∙∙ PE⟦ x₂ ⟧ - PE⟦ congExpr f x ⟧ = cong _ (cast↓Inj {[T]} PE⟦ x ⟧) - - cong-flat : ∀ {A A₁ a₀ a₁ } → (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) - → PathExpr A₁ a₀ a₁ - - → FlatPath (TyAt [T] A) false (cast↓ [T] A (f a₀)) - (cast↓ [T] A (f a₁)) - cong-flat f reflExpr = [fp] - cong-flat f (atomExpr p) = fp[ cong _ p ] - cong-flat f (compExpr x x₁ x₂) = - cong-flat f x fp++ cong-flat f x₁ fp++ cong-flat f x₂ - cong-flat f (congExpr f₁ x) = cong-flat (f ∘ f₁) x - - flat⟦_⟧ : ∀ {A a₀ a₁} → PathExpr A a₀ a₁ - → FlatPath (TyAt [T] A) false (cast↓ [T] A a₀) (cast↓ [T] A a₁) - flat⟦ reflExpr ⟧ = [fp] - flat⟦ atomExpr p ⟧ = fp[ cong (cast↓ [T] _) p ] - flat⟦ compExpr x x₁ x₂ ⟧ = flat⟦ x ⟧ fp++ flat⟦ x₁ ⟧ fp++ flat⟦ x₂ ⟧ - flat⟦ congExpr f x ⟧ = cong-flat f x - - - cong-flat≡ : ∀ {A₁ A a₀ a₁} → (pe : PathExpr A₁ a₀ a₁) → - (f : lookupTₙ [T] A₁ → lookupTₙ [T] A) → - (λ i → cast↓ [T] A (f (cast↓Inj PE⟦ pe ⟧ i))) ≡ - FP⟦ cong-flat f pe ⟧ - cong-flat≡ reflExpr f = cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) - cong-flat≡ (atomExpr p) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ lUnit _ - cong-flat≡ (compExpr pe pe₁ pe₂) f = - (cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-∙∙ _ _ _) ∙∙ - cong-∙∙ ((cast↓ [T] _ ∘ f)) _ _ _ ∙∙ - (λ i → doubleCompPath-elim - (cong-flat≡ pe f i) - (cong-flat≡ pe₁ f i) - (cong-flat≡ pe₂ f i) i)) - ∙∙ cong (_∙ FP⟦ cong-flat f pe₂ ⟧) - (fp++∙ (cong-flat f pe) (cong-flat f pe₁)) - ∙∙ fp++∙ _ (cong-flat f pe₂) - cong-flat≡ (congExpr f₁ pe) f = - cong (cong (cast↓ [T] _ ∘ f)) (cast↓Inj-sec _) ∙ cong-flat≡ pe (f ∘ f₁) - - pe≡flat : ∀ {A a₀ a₁} → (pe : PathExpr A a₀ a₁) → - PE⟦ pe ⟧ ≡ FP⟦ flat⟦ pe ⟧ ⟧ - pe≡flat reflExpr = refl - pe≡flat (atomExpr p) = lUnit _ - pe≡flat (compExpr pe pe₁ pe₂) = - (λ i → doubleCompPath-elim - (pe≡flat pe i) - (pe≡flat pe₁ i) - (pe≡flat pe₂ i) i) - ∙∙ cong (_∙ FP⟦ flat⟦ pe₂ ⟧ ⟧) (fp++∙ flat⟦ pe ⟧ flat⟦ pe₁ ⟧) - ∙∙ fp++∙ _ flat⟦ pe₂ ⟧ - - pe≡flat (congExpr f pe) = cong-flat≡ pe f - - -module PathTrm (A B : Type ℓ) where - data PathTrm : Type ℓ where - reflTrm : PathTrm - atomTrm : A → PathTrm - compTrm : PathTrm → PathTrm → PathTrm → PathTrm - congTrm : B → PathTrm → PathTrm - - module showPathTrm (showA : A → String) (showB : B → String) where - showPT : PathTrm → String - showPT reflTrm = "refl" - showPT (atomTrm x) = showA x - showPT (compTrm x x₁ x₂) = - "(" <> showPT x <> "∙∙" <> showPT x₁ <> "∙∙" <> showPT x₂ <> ")" - showPT (congTrm x x₁) = - "(" <> showB x <> "⟪" <> showPT x₁ <> "⟫)" - - -module _ {ℓ ℓ'} - {A B : Type ℓ} - {A' B' : Type ℓ'} - (f : A → R.TC A') - (g : B → R.TC B') where - open PathTrm - mmapPT : PathTrm A B → R.TC $ PathTrm A' B' - mmapPT reflTrm = pure reflTrm - mmapPT (atomTrm x) = ⦇ atomTrm (f x) ⦈ - mmapPT (compTrm x x₁ x₂) = - ⦇ compTrm (mmapPT x) (mmapPT x₁) (mmapPT x₂) ⦈ - mmapPT (congTrm x x₁) = - ⦇ congTrm (g x) (mmapPT x₁) ⦈ - - -module RTrm = PathTrm R.Term R.Term -module RTrm' = PathTrm (ℕ × R.Term) (ℕ × R.Term) -module StrTrm = PathTrm String String - -showRTrmTC : RTrm.PathTrm → R.TC String -showRTrmTC = - mmapPT - (R.formatErrorParts ∘ [_]ₑ) (R.formatErrorParts ∘ [_]ₑ) - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -showRTrmTC' : RTrm'.PathTrm → R.TC String -showRTrmTC' = - let q = λ (k , t) → - R.formatErrorParts (primShowNat k <> " " ∷ₑ [ t ]ₑ) - in mmapPT - q q - >=> (pure ∘ StrTrm.showPathTrm.showPT (idfun _) (idfun _) ) - -module _ ([T] : [Typeₙ]) where - reflExpr' : ∀ A (x : TyAt [T] A) → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A x) - reflExpr' A x = PE.reflExpr {[T] = [T]} {A} {cast↑ [T] A x} - - - atomExpr' : ∀ A {x y} → (p : x ≡ y) → - PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - atomExpr' A p = PE.atomExpr {[T] = [T]} {A} (cong (cast↑ [T] A) p) - - compExpr' : ∀ A {x y z w} - → PE.PathExpr [T] A x y → PE.PathExpr [T] A y z → PE.PathExpr [T] A z w - → PE.PathExpr [T] A x w - compExpr' A = PE.compExpr {[T] = [T]} {A = A} - - congExpr' : ∀ A A' {x y} → (f : TyAt [T] A → TyAt [T] A') - → PE.PathExpr [T] A (cast↑ [T] A x) (cast↑ [T] A y) - → PE.PathExpr [T] A' (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A x)))) - (cast↑ [T] A' (f (cast↓ [T] A (cast↑ [T] A y)))) - congExpr' A A' f x = PE.congExpr {[T] = [T]} {A = A} {A'} - (cast↑ [T] A' ∘ f ∘ cast↓ [T] A) x - -getEnd : ∀ {x y : A} → x ≡ y → A -getEnd p = p i0 - -module tryPathE --([T] : [Typeₙ]) - (TC[T]trm : R.Term) - ([TC[T]trm] [fns] : List R.Term) where - - inferA : R.Term → R.TC ℕ - inferA x = tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - - inferA' : R.Term → R.TC R.Term - inferA' = inferA >=> R.quoteTC - - - - try≡ : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - - - tryRefl : R.Term → R.TC (RTrm.PathTrm × R.Term) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - let x₀ = R.def (quote getEnd) v[ t ] - A ← inferA' t - ⦇ (RTrm.reflTrm , R.def (quote reflExpr') - (TC[T]trm v∷ A v∷ x₀ v∷ [])) ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - (t1 , t1') ← (try≡ k t1) - (t2 , t2') ← (try≡ k t2) - (t3 , t3') ← (try≡ k t3) - A ← inferA' t - pure (RTrm.compTrm t1 t2 t3 , - R.def (quote compExpr') - (TC[T]trm v∷ A v∷ t1' v∷ t2' v∷ t3' v∷ [])) - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm.PathTrm × R.Term) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t*) ← h tm - A ← inferA' t* - A' ← inferA' t - (t0 , t0') ← (try≡ k t*) - pure (RTrm.congTrm f t0 , - R.def (quote congExpr') - (TC[T]trm v∷ A v∷ A' v∷ f v∷ t0' v∷ []))) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm.PathTrm × R.Term) - atom x = do - A ← inferA' x - pure (RTrm.atomTrm x , R.def (quote atomExpr') - (TC[T]trm v∷ A v∷ x v∷ [])) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - -module tryTermE ([TC[T]trm] [fns] : List R.Term) where - - try≡ : ℕ → R.Term → R.TC (RTrm'.PathTrm) - - - tryRefl : R.Term → R.TC (RTrm'.PathTrm) - tryRefl t = do - _ ← R.noConstraints $ R.checkType - (R.con (quote reflCase) []) - (R.def (quote PathCase) ([ varg t ])) - ⦇ RTrm'.reflTrm ⦈ - - tryComp : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryComp zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryComp (suc k) t = do - tm ← R.noConstraints $ R.checkType - (R.con (quote compCase) (R.unknown v∷ R.unknown v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (t1 , t2 , t3) ← h tm - ⦇ RTrm'.compTrm (try≡ k t1) (try≡ k t2) (try≡ k t3) ⦈ - - where - - h : R.Term → R.TC (R.Term × R.Term × R.Term) - h (R.con (quote compCase) l) = match3Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - tryCong : ℕ → R.Term → R.TC (RTrm'.PathTrm) - tryCong zero _ = R.typeError [ (R.strErr "Magic number to low") ] - tryCong (suc k) t = - tryAllTC (R.typeError [ "not cong case" ]ₑ) - (zipWithIndex [fns]) - (λ (m , f) → do - tm ← R.noConstraints $ R.checkType - (R.con (quote congCase) (f v∷ [ varg R.unknown ])) - (R.def (quote PathCase) ([ varg t ])) - (_ , t) ← h tm - ⦇ (RTrm'.congTrm (m , f)) (try≡ k t) ⦈) - - where - - h : R.Term → R.TC (R.Term × R.Term) - h (R.con (quote congCase) l) = match2Vargs l - h _ = R.typeError [ (R.strErr "tryCompFail-h") ] - - - - atom : R.Term → R.TC (RTrm'.PathTrm) - atom x = do - k ← tryAllTC (R.typeError [ "notRecignisedType" ]ₑ) - (zipWithIndex [TC[T]trm]) - λ (k , Ty) → - R.checkType x (R.def (quote Path) - (Ty v∷ R.unknown v∷ R.unknown v∷ []) ) - >> pure k - R.returnTC (RTrm'.atomTrm (k , x)) - - - try≡ zero _ = R.typeError [ (R.strErr "Magic number to low") ] - try≡ (suc k) t = - R.catchTC - (tryRefl t) - (R.catchTC (tryComp k t) - (R.catchTC (tryCong k t) (atom t))) - - - --- funTypeView : R.Type → R.TC (Maybe (R.Type × R.Type)) --- funTypeView = {!!} - -groupoidSolverTerm : Bool → R.Term → R.Term → R.TC (R.Term × List R.ErrorPart) -groupoidSolverTerm debugFlag inpF hole = do - - -- inp₀ ← wait-for-type inp₀' - - -- R.typeError (niceAtomList [) -- >>= wait-for-type - (A' , (t0' , t1')) ← R.inferType hole >>= (get-boundaryWithDom ) >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary" ]) - (λ x → R.returnTC x) - - -- A ← wait-for-type A' - -- t0 ← wait-for-type t0' - -- t1 ← wait-for-type t1' - - - - let t0 = t0' - let t1 = t1' - - (AA , _) ← get-boundaryWithDom A' >>= Mb.rec - (R.typeError [ R.strErr "unable to get boundary2" ]) - (λ x → R.returnTC x) - inp₀ ← ([reflectedTy]→reflected[Typeₙ] <$> inferTypesFromFns inpF) >>= wait-for-term - - -- (R.typeError [ inp₀ ]ₑ) - - let inp = (R.con (quote _∷Tω_) (AA v∷ inp₀ v∷ [])) - -- (R.typeError [ inp ]ₑ) - - - - [t] ← reflected[Typeₙ]→[reflectedTy] inp - [f] ← reflected[Fns]→[reflectedFn] inpF - - (r0 , r0') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t0) - (r1 , r1') ← R.runSpeculative ((_, false) <$> tryPathE.try≡ inp [t] [f] 100 t1) - - - - -- (aL' , (_ , e0)) ← unMapAtoms [] r0' - -- (aL , (_ , e1)) ← unMapAtoms aL' r1' - -- let (_ , e0deas) = (Pℕ.⟦ e0 ⟧da∘r) - -- let (_ , e1deas) = (Pℕ.⟦ e1 ⟧da∘r) - show0 ← showRTrmTC r0 - show1 ← showRTrmTC r1 - - red0 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r0' v∷ [])]) >>= redList' - red1 ← (R.normalise $ R.def (quote fpF→T) v[ R.def (quote PE.flat⟦_⟧) (inp v∷ r1' v∷ [])]) >>= redList' - - - let invPa0 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red0 - invPa1 = Li.map - (λ t' → just (R.def (quote fpT≡F) (t' v∷ []))) - red1 - - let dbgInfo = --("r0Ty: ") ∷ₑ r0'Ty ∷nl - ("LHS: ") ∷ₑ show0 ∷nl - ("RHS: ") ∷ₑ show1 ∷nl - (niceAtomList red0 ++ ("" ∷nl niceAtomList red1)) - -- ∷ (R.strErr "RHS: ") ∷ (R.termErr $ t1) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPT (e0)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPT (e1)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTD (e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTD (e1deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "LHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e0deas)) - -- ∷ (R.strErr "\n") - -- ∷ (R.strErr "RHS: ") ∷ (R.strErr $ PℕS.showIPTI (ℕDeas→ℕInvol e1deas)) - - -- ∷ (R.strErr "\n") - -- (niceAtomList aL) - - let finalTrm0 = - just (R.def (quote PE.pe≡flat) (inp v∷ r0' v∷ [])) - ∷ invPa0 - - finalTrm1 = - just (R.def (quote PE.pe≡flat) (inp v∷ r1' v∷ [])) - ∷ invPa1 - - let finalTrm = fromMaybe (R.def (quote refl) []) $ foldPathTerms - (finalTrm0 ++ symPathTerms finalTrm1) - - pure (finalTrm , dbgInfo) - - -groupoidSolverMain : Bool → R.Term → R.Term → R.TC Unit -groupoidSolverMain debugFlag inpF' hole = do - inpF ← castTo[Fns] inpF' - inp ← inferTypesFromFns inpF - R.commitTC - ty ← R.withNormalisation true $ R.inferType hole >>= wait-for-term - hole' ← R.withNormalisation true $ R.checkType hole ty - - (solution , msg) ← groupoidSolverTerm debugFlag inpF hole' - R.catchTC - (R.unify hole solution) - (R.typeError msg) - -squareSolverMain : Bool → R.Term → R.Term → R.TC Unit -squareSolverMain debugFlag inpF' hole = do - inpF ← castTo[Fns] inpF' - inp ← inferTypesFromFns inpF - ty ← R.inferType hole >>= wait-for-term - hole' ← R.checkType (R.def (quote compPathR→PathP∙∙) (R.unknown v∷ [])) ty >>= extractMeta - >>= R.normalise - (solution , msg) ← groupoidSolverTerm debugFlag inpF hole' - R.catchTC - (R.unify hole' solution) - (R.typeError msg) - - R.catchTC - (R.unify hole (R.def (quote compPathR→PathP∙∙) (hole' v∷ []))) - (R.typeError [ R.strErr "xxx" ] ) - where - extractMeta : R.Term → R.TC R.Term - extractMeta (R.def _ tl) = matchVarg tl - extractMeta t = - R.typeError (R.strErr "extractMetaFail :" ∷ [ R.termErr t ]) - -macro - - ≡! : R.Term → R.TC Unit - ≡! = groupoidSolverMain true (R.con (quote [fn]) []) - - ≡!cong : R.Term → R.Term → R.TC Unit - ≡!cong = groupoidSolverMain true - - - sq! : R.Term → R.TC Unit - sq! = squareSolverMain false (R.con (quote [fn]) []) - - sq!cong : R.Term → R.Term → R.TC Unit - sq!cong = squareSolverMain true From ce967584d91663cc9463ef984fc0edcdf91da39f Mon Sep 17 00:00:00 2001 From: Marcin Grzybowski Date: Mon, 2 Sep 2024 23:34:53 +0200 Subject: [PATCH 20/20] examples tweaks --- Cubical/Tactics/GroupSolver/Example.agda | 4 +++- Cubical/Tactics/WildCatSolver/Example.agda | 21 +++++---------------- 2 files changed, 8 insertions(+), 17 deletions(-) diff --git a/Cubical/Tactics/GroupSolver/Example.agda b/Cubical/Tactics/GroupSolver/Example.agda index 66f6724c30..fb8d6bfbeb 100644 --- a/Cubical/Tactics/GroupSolver/Example.agda +++ b/Cubical/Tactics/GroupSolver/Example.agda @@ -67,7 +67,7 @@ module ℤexamples where open GroupStr (snd ℤGroup) - module _ k ([_]ᶻ : ℕ → fst ℤGroup) where + module _ k ([_]ᶻ : ℕ → fst ℤGroup) (someℤHom[_] : ℕ → GroupHom ℤGroup ℤGroup) where open Group-Solver ℓ-zero @@ -76,6 +76,8 @@ module ℤexamples where lhs rhs : (fst ℤGroup) lhs = fst (ℤHom k) ([ 1 ]ᶻ · [ 3 ]ᶻ) + · fst someℤHom[ 2 ] ([ 4 ]ᶻ · fst (someℤHom[ 1 ]) (inv [ 4 ]ᶻ · [ 4 ]ᶻ)) + · inv (fst someℤHom[ 2 ] ([ 4 ]ᶻ)) rhs = fst (ℤHom k) [ 1 ]ᶻ · fst (ℤHom k) [ 3 ]ᶻ lhs≡rhs : lhs ≡ rhs diff --git a/Cubical/Tactics/WildCatSolver/Example.agda b/Cubical/Tactics/WildCatSolver/Example.agda index ef33a19ef0..303b406bc7 100644 --- a/Cubical/Tactics/WildCatSolver/Example.agda +++ b/Cubical/Tactics/WildCatSolver/Example.agda @@ -50,7 +50,7 @@ module exampleC ℓ ℓ' where open Cat-Solver ℓ ℓ' - module _ (C C* : Category ℓ ℓ') (F : Functor C* C) where + module _ (C C* : Category ℓ ℓ') (F : Functor C* C) (G : Functor C C) where open Category C @@ -82,21 +82,10 @@ module exampleC ℓ ℓ' where (s : Hom[ F ⟅ w ⟆ , v ]) where - - - - -- pA pB : Hom[ F ⟅ y ⟆ , F ⟅ w ⟆ ] - -- pA = F ⟪ q *.⋆ r ⟫ - -- pB = F ⟪ q ⟫ ⋆ F ⟪ r ⟫ - - -- pA=pB : pA ≡ pB - -- pA=pB = solveCat (C ∷ C* ∷ []) - - - pA pB pC : Hom[ x , v ] - pA = (p ⋆ (id ⋆ F ⟪ q ⟫)) ⋆ (F ⟪ r ⟫ ⋆ s) - pB = ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⋆ (( F ⟪ r ⟫ ⋆ id) ⋆ s) - pC = p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) + pA pB pC : Hom[ G ⟅ x ⟆ , G ⟅ v ⟆ ] + pA = G ⟪ (p ⋆ (id ⋆ F ⟪ q ⟫)) ⟫ ⋆ G ⟪ (F ⟪ r ⟫ ⋆ s) ⟫ + pB = G ⟪ ((p ⋆ F ⟪ q *.⋆ (*.id *.⋆ *.id) ⟫ ) ⋆ F ⟪ *.id *.⋆ *.id ⟫) ⟫ ⋆ G ⟪ (( F ⟪ r ⟫ ⋆ id) ⋆ s) ⟫ + pC = id ⋆ G ⟪ p ⋆ (F ⟪ q *.⋆ r ⟫ ⋆ s) ⟫ ⋆ id pA=pB : pA ≡ pB pA=pB = solveCat