diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index 0693ad8b64..ff2ea413eb 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -80,7 +80,7 @@ jobs: - name: Download and install Agda from github if: steps.cache-external-restore.outputs.cache-hit != 'true' run: | - git clone https://github.com/agda/agda --branch ${{ env.AGDA_BRANCH }} --depth=1 + git clone https://github.com/${{ env.AGDA_REPO }}/agda --branch ${{ env.AGDA_BRANCH }} --depth=1 cd agda ${{ env.CABAL_INSTALL }} cd .. diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index 7a9e74119c..518bc7717d 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -1,20 +1,21 @@ module Cubical.Data.List.Properties where -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.GroupoidLaws + open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism - open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat +open import Cubical.Data.Maybe +open import Cubical.Data.Bool open import Cubical.Data.Sigma open import Cubical.Data.Sum as ⊎ hiding (map) open import Cubical.Data.Unit -open import Cubical.Data.List.Base as List - open import Cubical.Relation.Nullary +open import Cubical.Data.List.Base as List + module _ {ℓ} {A : Type ℓ} where ++-unit-r : (xs : List A) → xs ++ [] ≡ xs @@ -41,7 +42,7 @@ module _ {ℓ} {A : Type ℓ} where rev-rev-snoc : (xs : List A) (y : A) → Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl - rev-rev-snoc [] y = sym (lUnit refl) + rev-rev-snoc [] y = sym (compPath-filler refl refl) rev-rev-snoc (x ∷ xs) y i j = hcomp (λ k → λ @@ -120,7 +121,7 @@ private variable ℓ ℓ' : Level A : Type ℓ - B : Type ℓ' + B C : Type ℓ' x y : A xs ys : List A @@ -136,6 +137,14 @@ private safe-tail [] = [] safe-tail (_ ∷ xs) = xs +SafeHead' : List A → Type _ +SafeHead' [] = Unit* +SafeHead' {A = A} (x ∷ x₁) = A + +safeHead' : ∀ xs → SafeHead' {A = A} xs +safeHead' [] = tt* +safeHead' (x ∷ _) = x + cons-inj₁ : x ∷ xs ≡ y ∷ ys → x ≡ y cons-inj₁ {x = x} p = cong (safe-head x) p @@ -193,6 +202,15 @@ length-map : (f : A → B) → (as : List A) length-map f [] = refl length-map f (a ∷ as) = cong suc (length-map f as) +intersperse : A → List A → List A +intersperse _ [] = [] +intersperse a (x ∷ []) = x ∷ [] +intersperse a (x ∷ xs) = x ∷ a ∷ intersperse a xs + +join : List (List A) → List A +join [] = [] +join (x ∷ xs) = x ++ join xs + map++ : (f : A → B) → (as bs : List A) → map f as ++ map f bs ≡ map f (as ++ bs) map++ f [] bs = refl @@ -301,9 +319,13 @@ split++ (x₁ ∷ xs') ys' (x₂ ∷ xs) ys x = in zs , ⊎.map (map-fst (λ q i → p i ∷ q i)) (map-fst (λ q i → p (~ i) ∷ q i)) q -rot : List A → List A -rot [] = [] -rot (x ∷ xs) = xs ∷ʳ x +zipWithIndex : List A → List (ℕ × A) +zipWithIndex [] = [] +zipWithIndex (x ∷ xs) = (zero , x) ∷ map (map-fst suc) (zipWithIndex xs) + +repeat : ℕ → A → List A +repeat zero _ = [] +repeat (suc k) x = x ∷ repeat k x take[] : ∀ n → take {A = A} n [] ≡ [] take[] zero = refl @@ -318,6 +340,42 @@ lookupAlways a [] _ = a lookupAlways _ (x ∷ _) zero = x lookupAlways a (x ∷ xs) (suc k) = lookupAlways a xs k +lookupMb : List A → ℕ → Maybe A +lookupMb = lookupAlways nothing ∘S map just + +rot : List A → List A +rot [] = [] +rot (x ∷ xs) = xs ∷ʳ x + +rotN : ℕ → List A → List A +rotN n = iter n rot + +offset : A → ℕ → List A → List A +offset a n xs = repeat (substLen n xs) a ++ xs + where + substLen : ℕ → List A → ℕ + substLen zero _ = zero + substLen k@(suc _) [] = k + substLen (suc k) (_ ∷ xs) = substLen k xs + +offsetR : A → ℕ → List A → List A +offsetR a zero xs = xs +offsetR a (suc n) [] = repeat (suc n) a +offsetR a (suc n) (x ∷ xs) = x ∷ offsetR a n xs + +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 + + module List₂ where open import Cubical.HITs.SetTruncation renaming (rec to rec₂ ; map to map₂ ; elim to elim₂ ) @@ -346,3 +404,100 @@ module List₂ where List-comm-∥∥₂ : ∀ {ℓ} → List {ℓ} ∘ ∥_∥₂ ≡ ∥_∥₂ ∘ List List-comm-∥∥₂ = funExt λ A → isoToPath (Iso∥List∥₂List∥∥₂ {A = A}) + +zipWith : (A → B → C) → List A → List B → List C +zipWith f [] ys = [] +zipWith f (x ∷ xs) [] = [] +zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys + +alwaysZipWith : (Maybe A → Maybe B → C) → List A → List B → List C +alwaysZipWith f [] [] = [] +alwaysZipWith f [] ys = map (f nothing ∘ just) ys +alwaysZipWith f xs@(_ ∷ _) [] = map (flip f nothing ∘ just) xs +alwaysZipWith f (x ∷ xs) (y ∷ ys) = f (just x) (just y) ∷ alwaysZipWith f xs ys + +range : ℕ → List ℕ +range zero = [] +range (suc x) = x ∷ range x + +insertAt : ℕ → A → List A → List A +insertAt k a xs = take k xs ++ a ∷ drop k xs + +replaceAt : ℕ → A → List A → List A +replaceAt k a xs = take k xs ++ a ∷ drop (suc k) xs + +dropAt : ℕ → List A → List A +dropAt k xs = take k xs ++ drop (suc k) xs + + +findBy : (A → Bool) → List A → Maybe A +findBy t [] = nothing +findBy t (x ∷ xs) = if t x then just x else findBy t xs + + +catMaybes : List (Maybe A) → List A +catMaybes [] = [] +catMaybes (nothing ∷ xs) = catMaybes xs +catMaybes (just x ∷ xs) = x ∷ catMaybes xs + +fromAllMaybes : List (Maybe A) → Maybe (List A) +fromAllMaybes [] = just [] +fromAllMaybes (nothing ∷ xs) = nothing +fromAllMaybes (just x ∷ xs) = map-Maybe (x ∷_) (fromAllMaybes xs) + +maybeToList : Maybe A → List A +maybeToList nothing = [] +maybeToList (just x) = [ x ] + +listToMaybe : List A → Maybe A +listToMaybe [] = nothing +listToMaybe (x ∷ _) = just x + + +injectiveZipWith, : (xs ys : List (A × B)) → + map fst xs ≡ map fst ys → + map snd xs ≡ map snd ys → + xs ≡ ys +injectiveZipWith, [] [] x x₁ = refl +injectiveZipWith, [] (x₂ ∷ ys) x x₁ = ⊥.rec (¬nil≡cons x) +injectiveZipWith, (x₂ ∷ xs) [] x x₁ = ⊥.rec (¬nil≡cons (sym x)) +injectiveZipWith, (x₂ ∷ xs) (x₃ ∷ ys) x x₁ = + cong₂ _∷_ (cong₂ _,_ (cons-inj₁ x) (cons-inj₁ x₁)) + (injectiveZipWith, xs ys (cons-inj₂ x) (cons-inj₂ x₁)) + + +cart : List A → List B → List (A × B) +cart la lb = join (map (λ b → map (_, b) la) lb) + +filter : (A → Bool) → List A → List A +filter f [] = [] +filter f (x ∷ xs) = if f x then (x ∷ filter f xs) else (filter f xs) + + +module _ (_≟_ : Discrete A) where + nub : List A → List A + nub [] = [] + nub (x ∷ xs) = x ∷ filter (not ∘ Dec→Bool ∘ _≟ x) (nub xs) + + elem? : A → List A → Bool + elem? x [] = false + elem? x (x₁ ∷ x₂) = Dec→Bool (x ≟ x₁) or elem? x x₂ + + subs? : List A → List A → Bool + subs? xs xs' = foldr (_and_ ∘ flip elem? xs') true xs + +takeWhile : (A → Maybe B) → List A → List B +takeWhile f [] = [] +takeWhile f (x ∷ xs) with f x +... | nothing = [] +... | just y = y ∷ takeWhile f xs + +dropBy : (A → Bool) → List A → List A +dropBy _ [] = [] +dropBy f (x ∷ xs) = + if f x then (dropBy f xs) else (x ∷ xs) + +mapAt : (A → A) → ℕ → List A → List A +mapAt _ _ [] = [] +mapAt f (suc k) (x ∷ xs) = x ∷ mapAt f k xs +mapAt f zero (x ∷ xs) = f x ∷ xs diff --git a/Cubical/Data/Maybe/Properties.agda b/Cubical/Data/Maybe/Properties.agda index 981d2852fb..0bf9b6531b 100644 --- a/Cubical/Data/Maybe/Properties.agda +++ b/Cubical/Data/Maybe/Properties.agda @@ -178,3 +178,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/Foundations/Path.agda b/Cubical/Foundations/Path.agda index 0dd6bdb72e..056cb804d3 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Interpolate open import Cubical.Reflection.StrictEquiv @@ -159,6 +160,14 @@ flipSquare : {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} → Square a₀₋ a₁₋ a₋₀ a₋₁ → Square a₋₀ a₋₁ a₀₋ a₁₋ flipSquare sq i j = sq j i +flipSquareP : {A : I → I → Type ℓ} + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} {a₀₋ : PathP (λ j → A i0 j) a₀₀ a₀₁} + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} {a₁₋ : PathP (λ j → A i1 j) a₁₀ a₁₁} + {a₋₀ : PathP (λ i → A i i0) a₀₀ a₁₀} {a₋₁ : PathP (λ i → A i i1) a₀₁ a₁₁} + → SquareP A a₀₋ a₁₋ a₋₀ a₋₁ → SquareP (λ i j → A j i) a₋₀ a₋₁ a₀₋ a₁₋ +flipSquareP sq i j = sq j i + + module _ {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ ≡ a₁₁} {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} where @@ -211,13 +220,16 @@ sym≡flipSquare {x = x} P = sym (main refl P) -- Inverting both interval arguments of a square in Ω²A is the same as doing nothing sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) → P ≡ λ i j → P (~ i) (~ j) -sym-cong-sym≡id {x = x} P = sym (main refl P) - where - B : (q : x ≡ x) → I → Type _ - B q i = Path (x ≡ q i) (λ j → q (i ∨ ~ j)) λ j → q (i ∧ j) - - main : (q : x ≡ x) (p : refl ≡ q) → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p - main q = J (λ q p → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p) refl +sym-cong-sym≡id {x = x} P z i j = + hcomp (λ k → λ { + (i = i0) → P (~ k) (j ∨ z) + ;(i = i1) → x + ;(j = i0) → P (~ k ∧ ~ i) z + ;(j = i1) → x + ;(z = i0) → P (i ∨ ~ k) j + ;(z = i1) → P (~ i) (~ j) + }) + (P (~ i) (z ∧ ~ j)) -- Applying cong sym is the same as flipping a square in Ω²A flipSquare≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) @@ -447,8 +459,97 @@ Square→compPathΩ² {a = a} sq k i j = ; (k = i1) → cong (λ x → rUnit x r) (flipSquare sq) i j}) (sq j i) + +module _ {a b : A} {p q : a ≡ b} where + + cancel→pathsEq : p ∙ sym q ≡ refl → p ≡ q + cancel→pathsEq s i j = + hcomp (λ z → primPOr (~ i) (i ∨ j ∨ ~ j) + (λ _ → compPath-filler p (sym q) (~ z) j) λ _ → q (j ∧ z)) + (s i j) + +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' + + + pathFiber : {B : Type ℓ} (f : A → B) (b : B) {a a' : A} {t : f a ≡ b} {t' : f a' ≡ b} → ((a , t) ≡ (a' , t' )) → Σ[ e ∈ a ≡ a' ] (t ≡ cong f e ∙ t') pathFiber {A} {B} f b {a} {a'} {t} {t'} e = J (λ X _ → Σ[ e ∈ a ≡ fst X ] (t ≡ cong f e ∙ (snd X))) (refl , lUnit t) e + +invSides-filler-rot : {x y : A} (p : x ≡ y) → + invSides-filler p p ≡ symP (invSides-filler (sym p) (sym p)) +invSides-filler-rot {A = A} {x = x} {y} p z i j = + hcomp (λ k → primPOr (i ∨ ~ i) (j ∨ ~ j) + (λ _ → p (interpolateI (interpolateI i j (~ j)) (~ k ∧ z) (k ∨ z))) + (λ _ → p (interpolateI (interpolateI j i (~ i)) (~ k ∧ z) (k ∨ z)))) + (p z) + +module CompSquares + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ ≡ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ ≡ a₀₁₁} + {a₀₋₀ : a₀₀₀ ≡ a₀₁₀} {a₀₋₁ : a₀₀₁ ≡ a₀₁₁} + (a₀₋₋ : 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₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ ≡ a₁₁₀} {a₋₁₁ : a₀₁₁ ≡ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + where + + compSquaresSides : ∀ i j → I → Partial (i ∨ ~ i ∨ j ∨ ~ j ) A + compSquaresSides i j k (i = i0) = a₋₀₋ k j + compSquaresSides i j k (i = i1) = a₋₁₋ k j + compSquaresSides i j k (j = i0) = a₋₋₀ k i + compSquaresSides i j k (j = i1) = a₋₋₁ k i + + + compSquares : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁ + compSquares i j = + hcomp (compSquaresSides i j) (a₀₋₋ i j) + + + module _ (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) where + + + + compSquaresPath→Cube : compSquares ≡ a₁₋₋ → + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + compSquaresPath→Cube p z i j = + hcomp + (λ k → primPOr _ (i ∨ ~ i ∨ j ∨ ~ j) + (primPOr (~ z) z + (λ _ → hfill (compSquaresSides i j) (inS (a₀₋₋ i j)) (~ k)) (λ _ → p k i j)) + (compSquaresSides i j (z ∨ ~ k))) + (p i0 i j) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 260be6bfb9..63f70c4752 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -29,7 +29,7 @@ infixr 30 _∙₂_ infix 3 _∎ infixr 2 step-≡ _≡⟨⟩_ infixr 2.5 _≡⟨_⟩≡⟨_⟩_ -infixl 4 _≡$_ _≡$S_ +infixl 4 _≡$_ _≡$S_ _≡$≡_ _≡$'≡_ -- Basic theory about paths. These proofs should typically be -- inlined. This module also makes equational reasoning work with @@ -65,6 +65,8 @@ cong : (f : (a : A) → B a) (p : x ≡ y) → cong f p i = f (p i) {-# INLINE cong #-} +[_]$≡_ = cong + {- `S` stands for simply typed. Using `congS` instead of `cong` can help Agda to solve metavariables that may otherwise remain unsolved. -} @@ -371,6 +373,18 @@ funExtS⁻ eq x i = eq i x _≡$S_ = funExtS⁻ +_≡$≡_ : {B : I → Type ℓ'} + {f : (x : A) → B i0} {g : (x : A) → B i1} + → PathP (λ i → (x : A) → B i) f g + → {x y : A} → (x ≡ y) → PathP (λ i → B i) (f x) (g y) +(p ≡$≡ q) i = p i (q i) + +_≡$'≡_ : {B : I → Type ℓ'} + {f : {x : A} → B i0} {g : {x : A} → B i1} + → PathP (λ i → {x : A} → B i) (λ {x} → f {x}) (λ {x} → g {x}) + → {x y : A} → (x ≡ y) → PathP (λ i → B i) (f {x}) (g {y}) +(p ≡$'≡ q) i = p i {q i} + -- J for paths and its computation rule module _ (P : ∀ y → x ≡ y → Type ℓ') (d : P x refl) where diff --git a/Cubical/Foundations/Transport.agda b/Cubical/Foundations/Transport.agda index e044ba5998..3bd88d91b3 100644 --- a/Cubical/Foundations/Transport.agda +++ b/Cubical/Foundations/Transport.agda @@ -222,3 +222,22 @@ transport-filler-ua {A = A} {B = B} (e , _) a j i = in glue (λ { (i = i0) → a ; (i = i1) → tr j }) (hcomp (λ k → λ { (i = i0) → b ; (i = i1) → tr (j ∧ k) ; (j = i1) → tr (~ i ∨ k) }) (hcomp (λ k → λ { (i = i0) → tr (j ∨ k) ; (i = i1) → z ; (j = i1) → z }) z)) + +subst₂-≡ : ∀ {ℓ ℓ' ℓ''} → {A : Type ℓ} {B : Type ℓ'} (C : A → B → Type ℓ'') + → {a₀ a₁ : A} → (p : a₀ ≡ a₁) → {b₀ b₁ : B} → (q : b₀ ≡ b₁) + → ∀ c₀ c₁ + → (PathP (λ i → C (p i) (q i)) c₀ c₁) → + (transport (λ i → C a₀ (q i)) c₀ + ≡ transport (λ i → C (p (~ i)) b₁) c₁) +subst₂-≡ {B = B} C {a₀} = + J (λ a₁ p → {b₀ b₁ : B} → (q : b₀ ≡ b₁) + → ∀ c₀ c₁ + → (PathP (λ i → C (p i) (q i)) c₀ c₁) + → (transport (λ i → C a₀ (q i)) c₀ + ≡ transport (λ i → C (p (~ i)) b₁) c₁)) + λ {b₀} → J (λ b₁ q → + ∀ c₀ c₁ + → (PathP (λ i → C a₀ (q i)) c₀ c₁) + → (transport (λ i → C a₀ (q i)) c₀ + ≡ transport (λ i → C a₀ b₁) c₁)) + λ _ _ → transportRefl _ ∙∙_∙∙ sym (transportRefl _) diff --git a/Cubical/HITs/S2/Properties.agda b/Cubical/HITs/S2/Properties.agda index dac12b421a..93c2160cef 100644 --- a/Cubical/HITs/S2/Properties.agda +++ b/Cubical/HITs/S2/Properties.agda @@ -119,7 +119,8 @@ toSuspPresInvS² (surf i j) k r = l2 i j k = l1 j i (~ k) sym≡cong-sym-refl : ∀ {ℓ} {A : Type ℓ} {x : A} → sym≡cong-sym (λ _ _ → x) ≡ refl - sym≡cong-sym-refl {x = x} = transportRefl (λ _ _ _ → x) + sym≡cong-sym-refl {x = x} i j u v = + hcomp {φ = i ∨ j ∨ ~ j ∨ u ∨ ~ u ∨ v ∨ ~ v} (λ _ _ → x) x l1≡l2 : l1 ≡ l2 l1≡l2 = sym (sym≡flipSquare (λ i j → l1 j i)) diff --git a/Cubical/Reflection/Base.agda b/Cubical/Reflection/Base.agda index f09161db0f..9a31beb30e 100644 --- a/Cubical/Reflection/Base.agda +++ b/Cubical/Reflection/Base.agda @@ -7,19 +7,30 @@ 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 +open import Cubical.Reflection.Sugar.Base public + import Agda.Builtin.Reflection as R open import Agda.Builtin.String -_>>=_ = R.bindTC -_<|>_ = R.catchTC +instance + RawApplicativeTC : RawApplicative R.TC + RawApplicative._<$>_ RawApplicativeTC f x = R.bindTC x λ y → R.returnTC (f y) + RawApplicative.pure RawApplicativeTC = R.returnTC + RawApplicative._<*>_ RawApplicativeTC f x = R.bindTC f λ f → R.bindTC x λ x → R.returnTC (f x) -_>>_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → R.TC A → R.TC B → R.TC B -f >> g = f >>= λ _ → g +instance + RawMonadTC : RawMonad R.TC + RawMonad._>>=_ RawMonadTC = R.bindTC + RawMonad._>>_ RawMonadTC x y = R.bindTC x (λ _ → y) -infixl 4 _>>=_ _>>_ _<|>_ +instance + RawMonadFailTC : RawMonadFail R.TC (List R.ErrorPart) + RawMonadFail.fail RawMonadFailTC = R.typeError + RawMonadFail._<|>_ RawMonadFailTC = R.catchTC liftTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → R.TC A → R.TC B liftTC f ta = ta >>= λ a → R.returnTC (f a) @@ -32,6 +43,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/Reflection/Sugar.agda b/Cubical/Reflection/Sugar.agda new file mode 100644 index 0000000000..eca24b919c --- /dev/null +++ b/Cubical/Reflection/Sugar.agda @@ -0,0 +1,14 @@ +{- +This module provides essential definitions and the necessary machinery to enable syntactic sugar for monads and applicatives with instances. The content is split into two main submodules: + +1. **Base**: Contains fundamental definitions that only depend on `Agda.Primitive`, ensuring no dependencies on the Cubical library. This makes it universally usable across the library. +2. **Instances**: Provides instances of the defined classes for various entities from the Cubical library. + +The organization of this module emphasizes their utility as helper definitions for writing code and macros, rather than for formal reasoning about their categorical nature. +-} + +{-# OPTIONS --safe #-} +module Cubical.Reflection.Sugar where + +open import Cubical.Reflection.Sugar.Base public +open import Cubical.Reflection.Sugar.Instances public diff --git a/Cubical/Reflection/Sugar/Base.agda b/Cubical/Reflection/Sugar/Base.agda new file mode 100644 index 0000000000..208f7ccc9b --- /dev/null +++ b/Cubical/Reflection/Sugar/Base.agda @@ -0,0 +1,188 @@ +{- + +This module introduces several key classes and concepts for syntactic sugar in Agda: + +- **Raw Monad**: Defines a basic monadic structure. +- **Raw Monad Fail**: Extends `Raw Monad` to handle failure cases. +- **Raw Applicative**: Defines an applicative structure. +- **Raw Monad Transformer**: Facilitates monad transformers. +- **Identity Functor**: Provides an identity functor to use with the above structures. + +Key characteristics: + +- No dependencies on the Cubical library, ensuring broad usability. +- Designed to simplify macro writing, not for formal reasoning. +- Utilizes type-level `ω` for defining functors to ease instance writing without requiring universe polymorphism. + +All definitions necessary for syntactic sugar are implemented, allowing the use of multiple monads and applicatives correctly within the same context through instance resolution. + +-} + +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.Sugar.Base where + +open import Agda.Primitive public + using ( Level ) + renaming ( Set to Type + ; Setω to Typeω + ) + +private + variable + ℓ ℓ' : Level + A B C E : Type ℓ + +Functorω : Typeω +Functorω = ∀ {ℓ} → Type ℓ → Type ℓ + +private + variable + M : Functorω + +record IdentityF (A : Type ℓ) : Type ℓ where ; constructor identity ; field ; runIdentity : A + +open IdentityF public + + +record RawApplicative (F : Functorω) : Typeω where + field + _<$>_ : (A → B) → F A → F B + pure : A → F A + _<*>_ : F (A → B) → F A → F B + + infixl 4 _<*>_ _<$>_ + + + + +module _ (M : Functorω) {{RA : RawApplicative M}} where + + open RawApplicative RA public + + + + record RawMonad : Typeω where + field + _>>=_ : M A → (A → M B) → M B + _>>_ : M A → M B → M B + + infixl 4 _>>=_ _>>_ + + + + + module _ {{RM : RawMonad}} where + + open RawMonad RM public + + record RawMonadFail (E : Type ℓ') : Typeω where + field + fail : E → M A + _<|>_ : M A → M A → M A + + infixl 4 _<|>_ + + + module _ {{RME : RawMonadFail E}} where + open RawMonadFail RME public + + + +_>=&_ : ∀ {M : Functorω} {{_ : RawApplicative M}} → + (A → M B) → (B → C) → A → M C +(x >=& x₁) x₂ = x₁ <$> x x₂ + +joinM : ∀ {M : Functorω} {{_ : RawApplicative M}} {{_ : RawMonad M}} → + M (M A) → M A +joinM x = x >>= (λ x → x) + +_>=>_ : ∀ {M : Functorω} {{_ : RawApplicative M}} {{_ : RawMonad M}} → + (A → M B) → (B → M C) → A → M C +(x >=> x₁) x₂ = x x₂ >>= x₁ + +instance + RawApplicativeIdentityF : RawApplicative IdentityF + runIdentity ((RawApplicativeIdentityF RawApplicative.<$> f) x) = f (runIdentity x) + runIdentity (RawApplicative.pure RawApplicativeIdentityF x) = x + runIdentity ((RawApplicativeIdentityF RawApplicative.<*> f) x) = runIdentity f (runIdentity x) + + + RawMonadIdentityF : RawMonad IdentityF + (RawMonadIdentityF RawMonad.>>= x) y = y (runIdentity x) + (RawMonadIdentityF RawMonad.>> _) y = y + + +RawApplicativeIdentityM : RawApplicative (λ {_} X → X) +RawApplicative._<$>_ RawApplicativeIdentityM f = f +RawApplicative.pure RawApplicativeIdentityM x = x +RawApplicative._<*>_ RawApplicativeIdentityM f x = f x + + +RawMonadIdentityM : RawMonad (λ {_} X → X) {{RawApplicativeIdentityM}} +(RawMonadIdentityM RawMonad.>>= x) y = y x +(RawMonadIdentityM RawMonad.>> _) y = y + + +record RawMonadTransformer (T : Functorω → Functorω) : Typeω where + field + lifting : {F : Functorω} → {{RA : RawApplicative F}} → {{_ : RawMonad F}} → F A → T F A + + applicativeLiftT : {F : Functorω} → {{RA : RawApplicative F}} + → {{RM : RawMonad F}} → RawApplicative (T F) + + monadLiftT : {F : Functorω} → {{RA : RawApplicative F}} → {{_ : RawMonad F}} + → RawMonad (T F) {{applicativeLiftT}} + + +open RawMonadTransformer public + + +record [_RMT_]_ (T : Functorω → Functorω) (F : Functorω) {ℓ} (A : Type ℓ) : Type ℓ where + constructor wrap + field + unwrap : T F A + +open [_RMT_]_ public + + + + +module _ {T} {{rmt : RawMonadTransformer T}} where + + + private + instance + ApplicativeLiftT' : {F : Functorω} → + {{_ : RawApplicative F}} → {{_ : RawMonad F}} → RawApplicative (T F) + ApplicativeLiftT' = applicativeLiftT rmt + + instance + MonadLiftT' : {F : Functorω} → + {{_ : RawApplicative F}} → {{_ : RawMonad F}} → RawMonad (T F) + MonadLiftT' = monadLiftT rmt + + liftM : {F : Functorω} → {{_ : RawApplicative F}} → {{_ : RawMonad F}} → F A → [ T RMT F ] A + liftM x = wrap (lifting rmt x) + + instance + ApplicativeLiftT : {{RA : RawApplicative M}} + → {{RM : RawMonad M}} → RawApplicative ([ T RMT M ]_) + unwrap ((ApplicativeLiftT RawApplicative.<$> x) x₁) = x <$> unwrap x₁ + unwrap (RawApplicative.pure ApplicativeLiftT x) = pure x + unwrap ((ApplicativeLiftT RawApplicative.<*> x) x₁) = unwrap x <*> unwrap x₁ + + MonadLiftT : {{_ : RawApplicative M}} → {{_ : RawMonad M}} → RawMonad ([ T RMT M ]_) + unwrap ((MonadLiftT RawMonad.>>= x) x₁) = unwrap x >>= λ v → unwrap (x₁ v) + unwrap ((MonadLiftT RawMonad.>> x) x₁) = unwrap x >> unwrap x₁ + +_<|>>=_ : {M : Functorω} ⦃ RA : RawApplicative M ⦄ ⦃ RM : RawMonad M ⦄ + {ℓ : Level} {E : Type ℓ} ⦃ RME : RawMonadFail M E ⦄ + {ℓ' ℓ'' : Level} {A : Type ℓ'} {B : Type ℓ''} → + (A → M B) → (A → M B) → (A → M B) +(f <|>>= g) a = f a <|> g a + +_>>=<|>_ : {M : Functorω} ⦃ RA : RawApplicative M ⦄ ⦃ RM : RawMonad M ⦄ + {ℓ : Level} {E : Type ℓ} ⦃ RME : RawMonadFail M E ⦄ + {ℓ' ℓ'' : Level} {A : Type ℓ'} {B : Type ℓ''} → + (A → M B) → E → (A → M B) +(f >>=<|> e) x = f x <|> fail e diff --git a/Cubical/Reflection/Sugar/Instances.agda b/Cubical/Reflection/Sugar/Instances.agda new file mode 100644 index 0000000000..3513d4d4d4 --- /dev/null +++ b/Cubical/Reflection/Sugar/Instances.agda @@ -0,0 +1,161 @@ +{- + +This module provides specific instances for various classes defined in `Sugar.Base`, integrating definitions from the Cubical library. The instances cover several widely-used types and structures within the Cubical library. + +Key Features: + +- **Applicative and Monad Instances**: + - `Maybe`: Instance implementations for `RawApplicative` and `RawMonad`. + - `List`: Instance implementations for `RawApplicative` and `RawMonad`. + - `Sum`: Supports sum types (`⊎`), with instances for both `RawApplicative` and `RawMonad`. + +- **Monad Transformer Instances**: + - `State`: Provides a monad transformer instance for state computations. + - `Plus`: Provides a monad transformer instance for sum types. + +- **Utility Functions**: + - Functions like `mapM`, `concatMapM`, `foldlM`, and `foldrM` facilitate common monadic operations on lists. + - Functions `get` and `modify` support stateful computations. + + +Note how the use of type-level `ω` simplifies the writing of instances without dealing with universe polymorphism, wich is instead dealt with in de definition of functions relaying on those instances. + +-} + +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.Sugar.Instances where + + +open import Cubical.Foundations.Function + + +open import Cubical.Reflection.Sugar.Base +open import Cubical.Data.List as L +import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.Bool +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + + +module _ {M : Functorω} {{_ : RawApplicative M}} {{_ : RawMonad M}} where + + mapM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → M B) → List A → M (List B) + mapM f [] = ⦇ [] ⦈ + mapM f (x ∷ xs) = ⦇ f x ∷ mapM f xs ⦈ + + mapM-snd : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {A' : Type ℓ'} {B : Type ℓ''} + → (A → M B) → A' × A → M (A' × B) + mapM-snd f (x , y) = ⦇ ⦇ x ⦈ , f y ⦈ + + concatMapM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → M (List B)) → List A → M (List B) + concatMapM f [] = ⦇ [] ⦈ + concatMapM f (x ∷ xs) = ⦇ f x ++ concatMapM f xs ⦈ + + + foldlM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (B → A → M B) → B → List A → M B + foldlM f b [] = pure b + foldlM f b (x ∷ xs) = f b x >>= (flip (foldlM f)) xs + + + foldrM : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → B → M B) → B → List A → M B + foldrM f b [] = pure b + foldrM f b (x ∷ xs) = foldrM f b xs >>= f x + + +State₀T : Type → Functorω → Functorω +State₀T S F T = S → F (T × S) + + +get : {F : Functorω} {{FA : RawApplicative F }} {{FM : RawMonad F }} {S : Type} → [ State₀T S RMT F ] S +unwrap get s = pure (s , s) + +modify : {F : Functorω} {{FA : RawApplicative F }} {{FM : RawMonad F }} {S : Type} → + (S → S) → [ State₀T S RMT F ] Unit +unwrap (modify f) s = pure (_ , f s) + + +Plus₀T : Type → Functorω → Functorω +Plus₀T E F T = F (E ⊎.⊎ T) + +module _ where + open RawMonadTransformer + + instance + rawMonadTransformerState₀ : ∀ {S} → RawMonadTransformer (State₀T S) + (applicativeLiftT rawMonadTransformerState₀ RawApplicative.<$> x) = (map-fst x <$>_) ∘S_ + RawApplicative.pure (applicativeLiftT rawMonadTransformerState₀) x = pure ∘S (x ,_) + (applicativeLiftT rawMonadTransformerState₀ RawApplicative.<*> f) x s = + f s >>= uncurry (_<$>_ ∘S map-fst) ∘S map-snd x + + (monadLiftT rawMonadTransformerState₀ RawMonad.>>= x) y s = x s >>= uncurry y + + (monadLiftT rawMonadTransformerState₀ RawMonad.>> x) y s = x s >>= y ∘S snd + lifting rawMonadTransformerState₀ x s = (_, s) <$> x + + rawMonadTransformerPlus₀ : ∀ {S} → RawMonadTransformer (Plus₀T S) + (applicativeLiftT rawMonadTransformerPlus₀ RawApplicative.<$> x) = + (⊎.map (idfun _) x) <$>_ + RawApplicative.pure (applicativeLiftT rawMonadTransformerPlus₀) x = pure (⊎.inr x) + (applicativeLiftT rawMonadTransformerPlus₀ RawApplicative.<*> f) x = + f >>= ⊎.rec (pure ∘S ⊎.inl) λ f → ⊎.map (idfun _) f <$> x + (monadLiftT rawMonadTransformerPlus₀ RawMonad.>>= x) y = + x >>= ⊎.rec (pure ∘S ⊎.inl) y + (monadLiftT rawMonadTransformerPlus₀ RawMonad.>> x) y = + x >>= ⊎.rec (pure ∘S ⊎.inl) (const y) + lifting rawMonadTransformerPlus₀ x = ⊎.inr <$> x + + + ApplicativeSum : {E : Type} → RawApplicative (E ⊎.⊎_) + ApplicativeSum = applicativeLiftT rawMonadTransformerPlus₀ {{_}} {{RawMonadIdentityM}} + + MonadSum : {E : Type} → RawMonad (E ⊎.⊎_) + MonadSum = monadLiftT rawMonadTransformerPlus₀ {{_}} {{RawMonadIdentityM}} + + + +instance + ApplicativeMaybe : RawApplicative Maybe + RawApplicative._<$>_ ApplicativeMaybe = map-Maybe + RawApplicative.pure ApplicativeMaybe = just + (ApplicativeMaybe RawApplicative.<*> nothing) x₁ = nothing + (ApplicativeMaybe RawApplicative.<*> just x) nothing = nothing + (ApplicativeMaybe RawApplicative.<*> just x) (just x₁) = just (x x₁) + + + MonadMaybe : RawMonad Maybe + RawMonad._>>=_ MonadMaybe = flip (Mb.rec nothing) + RawMonad._>>_ MonadMaybe = flip (Mb.rec nothing ∘ const) + + ApplicativeList : RawApplicative List + RawApplicative._<$>_ ApplicativeList = L.map + RawApplicative.pure ApplicativeList = [_] + (ApplicativeList RawApplicative.<*> fs) xs = L.map (uncurry _$_) (cart fs xs) + + + MonadList : RawMonad List + RawMonad._>>=_ MonadList xs f = L.join (map f xs) + RawMonad._>>_ MonadList xs ys = L.join (map (λ _ → ys) xs) + + +when : ∀ {M : Functorω} {{_ : RawApplicative M}} → Bool → M Unit → M Unit +when {M} false x = pure _ +when {M} true x = x + + + +infixl 4 _<⊎>_ + + +private + variable + ℓ : Level + A B E : Type ℓ + +_<⊎>_ : {M : Functorω} {{_ : RawApplicative M}} {{_ : RawMonad M}} {{_ : RawMonadFail M E}} → + (M A) → (M B) → M (A ⊎.⊎ B) +a <⊎> b = (⊎.inl <$> a) <|> (⊎.inr <$> b) diff --git a/Cubical/Tactics/PathSolver/CartesianExperiment.agda b/Cubical/Tactics/PathSolver/CartesianExperiment.agda new file mode 100644 index 0000000000..fe8403db81 --- /dev/null +++ b/Cubical/Tactics/PathSolver/CartesianExperiment.agda @@ -0,0 +1,263 @@ +{- +This module serves as an experimental demonstration of the potential usage of cubical-flavored reflection machinery. It focuses on transforming terms to a form where all interval expressions are either `i0`, `i1`, or a single interval variable (This transformation excludes the implicit `φ` argument of `hcomp`, which is effectively a face expression), effectively mimicking the Cartesian cubical theory. + +### Example Transformations + +- **`cpfCC`**: Demonstrates the transformation on a compPath-filler. +- **`assocCC`**: Demonstrates the transformation on a paths associativity. +- ** `rot-refl'CC`**: Examples of transformed cubes. +-} + +{-# OPTIONS --safe #-} +module Cubical.Tactics.PathSolver.CartesianExperiment where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Interpolate + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sigma + + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar + +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.QuoteCubical +open import Cubical.Tactics.Reflection.CuTerm +open import Cubical.Tactics.Reflection.Error + + +private + variable + ℓ : Level + + +-- really just refl ∙_ +reComp : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → x ≡ y +reComp p i = + hcomp {φ = i ∨ ~ i} (λ k _ → (p (i ∧ k))) (p i0) + +-- really just lUnit +reFill : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → p ≡ reComp p +reFill p j i = + hcomp {φ = i ∨ ~ i ∨ ~ j} (λ k _ → (p (i ∧ k))) (p i0) + + +addConstSubfaces : ℕ → CuTermNC → R.TC CuTermNC +addConstSubfaces = h + where + + addMiss : ℕ → List (SubFace × CuTermNC) → CuTermNC → R.TC (List (SubFace × CuTermNC)) + addMiss dim xs xb = do + newSfs ← catMaybes <$> mapM mbTermForFace msf + pure (newSfs ++fe× xs) + where + msf = missingSubFaces dim (L.map fst xs) + + mbTermForFace : SubFace → R.TC (Maybe (SubFace × CuTermNC)) + mbTermForFace sf = do + cOnSF ← cuEvalN sf (hco xs xb) + if (allCellsConstant? (suc (sfDim sf)) cOnSF) + then pure $ just (sf , cell (liftVars (mostNestedCap cOnSF))) + else ⦇ nothing ⦈ + + h : ℕ → CuTermNC → R.TC CuTermNC + hh : List (SubFace × CuTermNC) → R.TC (List (SubFace × CuTermNC)) + + h dim (hco x x₁) = do + x' ← hh x + xb ← (h dim x₁) + ⦇ hco (addMiss dim x' xb) ⦇ xb ⦈ ⦈ + h dim (cell' x x₁) = pure $ cell' x x₁ + h dim (𝒄ong' x x₁) = R.typeError [ "notImplemented" ]ₑ + + hh [] = ⦇ [] ⦈ + hh ((sf , x) ∷ xs) = + ⦇ ⦇ ⦇ sf ⦈ , h (suc (sfDim sf)) x ⦈ ∷ (hh xs) ⦈ + + +module unConnect (do-fill : Bool) where + + unConnCell : ℕ → R.Term → R.Term → R.TC CuTermNC + unConnCell dim jT (R.var k (z₀ v∷ v[ z₁ ])) = + (if do-fill then (pure ∘S cell) + else (quoteCuTermNC nothing dim >=> addConstSubfaces dim)) + (R.def (quote reFill) + (vlam "𝒾" + ((R.def (quote reFill) (R.var (suc k) v[ 𝒗 zero ] v∷ (liftVars jT) v∷ v[ liftVars z₁ ]))) + v∷ jT v∷ v[ z₀ ])) + + unConnCell dim jT (R.var k v[ z ]) = + (if do-fill then (pure ∘S cell) else (quoteCuTermNC nothing dim)) + (R.def (quote reFill) ((R.var k []) v∷ jT v∷ v[ z ])) + unConnCell _ _ t = pure $ cell' _ t + + + unConnS : List (SubFace × CuTermNC) → R.TC (List (SubFace × CuTermNC)) + + unConnA : ℕ → List (Hco ⊥ Unit) → R.TC (List (Hco ⊥ Unit)) + + + unConn : ℕ → CuTermNC → R.TC (CuTermNC) + unConn dim (hco x x₁) = ⦇ hco (unConnS x) (unConn dim x₁) ⦈ + unConn dim (cell' x x₁) = + if do-fill + then unConnCell (suc dim) (𝒗 dim) (liftVarsFrom (suc zero) dim x₁) + else unConnCell dim (endTerm true) x₁ + unConn dim (𝒄ong' {cg = cg} x x₁) = 𝒄ong' {cg = cg} x <$> unConnA dim x₁ + + unConnS [] = ⦇ [] ⦈ + unConnS ((sf , x) ∷ xs) = ⦇ ⦇ ⦇ (sf ++ (if do-fill then [ nothing ] else [])) ⦈ + , unConn (suc (sfDim sf)) x ⦈ ∷ unConnS xs ⦈ + + unConnA _ [] = ⦇ [] ⦈ + unConnA dim (hcodata x x₁ ∷ xs) = ⦇ ⦇ hcodata (unConnS x) (unConn dim x₁) ⦈ ∷ (unConnA dim xs) ⦈ + + + +unConn = unConnect.unConn false +unConnFill = unConnect.unConn true + + +module _ (dim : ℕ) where + macro + unConnTest : R.Term → R.Term → R.TC Unit + unConnTest t h = do + cu ← (extractCuTermNC nothing dim t) + cu' ← unConn dim cu + te0 ← ppCT dim 100 cu + te0' ← ppCT dim 100 cu' + wrapError h $ + "input:" ∷nl (indentₑ 4 te0) + ++nl "\n∨,∧,~ - removed :" ∷nl (indentₑ 4 te0') + + unConnTest'' : R.Term → R.Term → R.TC Unit + unConnTest'' t h = do + cu ← (extractCuTermNC nothing dim t) + cu' ← unConn dim cu + te0 ← ppCT dim 100 cu + te0' ← ppCT dim 100 cu' + R.typeError te0' + + unConnM : R.Term → R.Term → R.TC Unit + unConnM t h = do + cu ← (extractCuTermNC nothing dim t) + cu' ← unConn dim cu + R.unify (toTerm dim cu') h + + unConnM≡ : R.Term → R.Term → R.TC Unit + unConnM≡ t h = do + cu ← (extractCuTermNC nothing dim t) + cu' ← unConnFill dim cu + let cu'T = toTerm (suc dim) cu' + -- cu'' ← R.checkType cu'T (R.def (quote PathP) (R.unknown v∷ R.unknown v∷ v[ R.unknown ])) + R.unify cu'T h + + + +module _ {A : Type ℓ} {x y z w : A} (p : x ≡ y)(q : y ≡ z)(r : z ≡ w) where + + cpfCC : ResultIs + ("input: " ∷ + " " ∷ + " 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₀ " ∷ + " ║ (𝓲₀=0) → q 𝓲₁ " ∷ + " ║ (𝓲₁=0) → p (~ 𝓲₀ ∨ ~ 𝒛₀) " ∷ + " ║ (𝓲₁=1) → r (𝓲₀ ∧ 𝒛₀) " ∷ + " ├─────────── " ∷ + " │ q 𝓲₁ " ∷ + " └─────────── " ∷ + "∨,∧,~ - removed : " ∷ + " " ∷ + " 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₀ " ∷ + " ║ (𝓲₀=0) → " ∷ + " ║ 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₁ " ∷ + " ║ ║ (𝓲₁=0) → y " ∷ + " ║ ║ (𝓲₁=1) → q 𝒛₁ " ∷ + " ║ ├─────────── " ∷ + " ║ │ y " ∷ + " ║ └─────────── " ∷ + " ║ (𝓲₁=0) → " ∷ + " ║ 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₁ " ∷ + " ║ ║ (𝒛₀=1)(𝓲₀=1) → x " ∷ + " ║ ║ (𝒛₀=0) → p 𝒛₁ " ∷ + " ║ ║ (𝓲₀=0) → p 𝒛₁ " ∷ + " ║ ├─────────── " ∷ + " ║ │ x " ∷ + " ║ └─────────── " ∷ + " ║ (𝓲₁=1) → " ∷ + " ║ 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₁ " ∷ + " ║ ║ (𝒛₀=0) → z " ∷ + " ║ ║ (𝓲₀=0) → z " ∷ + " ║ ║ (𝒛₀=1)(𝓲₀=1) → r 𝒛₁ " ∷ + " ║ ├─────────── " ∷ + " ║ │ z " ∷ + " ║ └─────────── " ∷ + " ├─────────── " ∷ + " │ " ∷ + " │ 𝒉𝒄𝒐𝒎𝒑 λ 𝒛₀ " ∷ + " │ ║ (𝓲₁=0) → y " ∷ + " │ ║ (𝓲₁=1) → q 𝒛₀ " ∷ + " │ ├─────────── " ∷ + " │ │ y " ∷ + " │ └─────────── " ∷ + " └─────────── " ∷ []) + cpfCC = unConnTest (suc (suc zero)) λ (i j : I) → doubleCompPath-filler p q r i j + + assocCC : Square _ _ _ _ + assocCC = unConnM (suc (suc zero)) λ (i j : I) → assoc p q r i j + + + +module Sq-rot-refl {A : Type ℓ} + {a : A} + (s : Square {a₀₀ = a} refl refl refl refl) where + + rot-refl : Cube + s (λ i j → s j (~ i)) + (λ i j → a) (λ i j → a) + (λ i j → a) (λ i j → a) + rot-refl k i j = + hcomp (λ l → λ { (i = i0) → s (~ l) (j ∨ k) + ; (i = i1) → a + ; (j = i0) → s (~ l) (~ i ∧ k) + ; (j = i1) → a + ; (k = i0) → s (i ∨ ~ l) j + ; (k = i1) → s (j ∨ ~ l) (~ i) + }) + a + + + + rot-refl' : s ≡ λ i j → s j (~ i) + rot-refl' t i j = + hcomp (λ l → λ { (t = i0) → s i j + ; (t = i1) → s j (~ i) + ; (i = i0) → s (~ l ∧ t ∧ j) ((~ t ∧ j) ∨ t ∨ j) + ; (i = i1) → s ((~ l ∧ ~ t ∨ (t ∧ j) ∨ j) ∨ l ∨ ~ t ∨ (t ∧ j) ∨ j) (~ t ∧ j) + ; (j = i0) → s (~ l ∧ ~ t ∧ i) (t ∧ ~ i) + ; (j = i1) → s ((~ l ∧ (~ t ∧ i) ∨ t ∨ i) ∨ l ∨ (~ t ∧ i) ∨ t ∨ i) + (~ t ∨ (t ∧ ~ i) ∨ ~ i) + }) + (s ((~ t ∧ i) ∨ (t ∧ j) ∨ i ∧ j) ((~ t ∧ j) ∨ (t ∧ ~ i) ∨ j ∧ ~ i)) + + + s' : Square {a₀₀ = a} refl refl refl refl + s' = unConnM (suc (suc zero)) (λ (i j : I) → s i j) + + s'-rot : Square {a₀₀ = a} refl refl refl refl + s'-rot = unConnM (suc (suc zero)) (λ (i j : I) → s j (~ i)) + + + rot-refl'CC : s' ≡ s'-rot + rot-refl'CC = unConnM (suc (suc (suc zero))) λ (z i j : I) → rot-refl' z i j + diff --git a/Cubical/Tactics/PathSolver/CongComp.agda b/Cubical/Tactics/PathSolver/CongComp.agda new file mode 100644 index 0000000000..c8408ab3c5 --- /dev/null +++ b/Cubical/Tactics/PathSolver/CongComp.agda @@ -0,0 +1,147 @@ +{- + +This module provides utilities for generalizing the `cong-∙` lemma from `GroupoidLaws.agda`. +The utilities handle multiple arguments (implicit and explicit) and multiple dimensions, applying lemma recursively to subterms at all depths. + + +- **`appCongs` Module**: Applies a function recursively over composition terms. +- **`fillCongs` Module**: Computes fillers between terms when the function is applied at various levels of the term structure. + +-} + +{-# OPTIONS --safe #-} +module Cubical.Tactics.PathSolver.CongComp where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sigma + + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.CuTerm +open import Cubical.Tactics.Reflection.Variables + + +getSide : ∀ {A} → SubFace → Hco A Unit → CuTerm' A Unit +getSide {A = A} sf (hcodata l y) = + Mb.rec + ((let msf : SubFace × CuTerm' A Unit → Maybe (SubFace × CuTerm' A Unit) + msf = λ (sf' , t) → + map-Maybe + (λ sf'' → (injSF sf sf'') , cuEval (nothing ∷ (injSF sf' sf'')) t) + (sf' sf∩ sf) + l' = L.catMaybes (L.map msf l) + y' = (cuEval sf y) + in cell (R.lit (R.string "todo in getSide")) ) + ) + ((λ (sf' , f) → f)) + (findBy (⊂⊃? ∘S (sf _) ∘S fst) l) + + + +module appCongs where + + appCongs : ℕ → ℕ → CuTerm → CuTermNC + cuCong1 : R.Term → CuTermNC → CuTermNC + cuCong1L : R.Term → List (SubFace × CuTermNC) → List (SubFace × CuTermNC) + cuCong1L t [] = [] + cuCong1L t ((sf , x) ∷ xs) = + (sf , cuCong1 (liftVarsFrom 1 1 (subfaceCell (nothing ∷ sf) t)) x) ∷ cuCong1L t xs + + cuCong1 t (hco x x₁) = hco (cuCong1L t x) (cuCong1 t x₁) + cuCong1 t (cell x₁) = cell (substTms [ x₁ ] t) + + congCus : ℕ → ℕ → R.Term → List (Hco ⊥ Unit) → CuTermNC + congCus zero _ _ _ = cell (R.lit (R.string "out of fuel in congCus")) + congCus (suc fuel) dim t ((hcodata x y) ∷ []) = cuCong1 t (hco x y) + congCus (suc fuel) dim t xs @(_ ∷ _ ∷ _) = --cell (R.lit (R.string "todo")) + let lid = appCongs fuel dim (𝒄ongF t (L.map (CuTermNC→CuTerm ∘S Hco.bottom) xs)) + + in hco (L.map ff sfUnion) lid + where + sfUnion = foldr (_++fe_ ∘S L.map fst ∘S Hco.sides) [] xs + + ff : SubFace → SubFace × CuTerm' ⊥ Unit + ff sf = sf , + let ts = L.map (getSide sf) xs + in appCongs fuel (suc (sfDim sf)) + (𝒄ongF + (liftVarsFrom 1 (length xs) + (subfaceCell ((repeat (length xs) nothing) ++ sf) t)) (L.map CuTermNC→CuTerm ts)) + + congCus _ _ t [] = cell t + + + + appCongsS : ℕ → List (SubFace × CuTerm) → List (SubFace × CuTermNC) + appCongsS zero _ = [] + appCongsS _ [] = [] + appCongsS (suc fuel) ((sf , x) ∷ xs) = + (sf , appCongs fuel (suc (sfDim sf)) x) ∷ appCongsS fuel xs + + appCongs zero _ _ = cell (R.lit (R.string "out of fuel in normal𝑪ong")) + appCongs (suc fuel) dim (hco x x₁) = + hco (appCongsS fuel x) (appCongs fuel dim x₁) + appCongs _ dim (cell' x x₁) = cell' x x₁ + appCongs (suc fuel) dim (𝒄ong' x x₁) = + congCus fuel dim x + (L.map (λ (hcodata sides bottom) → + hcodata (appCongsS fuel sides) (appCongs fuel dim bottom)) x₁) + +appCongs : ℕ → CuTerm → CuTermNC +appCongs = appCongs.appCongs 100 + + +module fillCongs where + + + fillCongs : ℕ → ℕ → CuTerm → CuTermNC + congFill : ℕ → ℕ → R.Term → List (Hco Unit Unit) → CuTermNC + congFill fuel dim t xs = + let lid = fillCongs fuel dim $ 𝒄ongF t (L.map (Hco.bottom) xs) + in hco (((repeat dim nothing ∷ʳ just false) , f0) ∷ + L.map ff sfUnion) lid + where + sfUnion = foldr (_++fe_ ∘S L.map fst ∘S Hco.sides) [] xs + + ff : SubFace → SubFace × CuTermNC + ff sf = sf , + let ts : List (CuTerm) + ts = L.map ((getSide sf)) xs + in fillCongs fuel (suc (sfDim sf)) $ + 𝒄ongF ((liftVarsFrom 1 (length xs) + (subfaceCell ((repeat (length xs) nothing) ++ sf) t))) ts + + f0 : CuTermNC + f0 = cell' _ (substTms ( + L.map ((ToTerm.toTermFill {Unit} {Unit} (defaultCtx dim))) + xs --xs + ) (liftVarsFrom 1 (length xs) t)) + + fillCongsS : ℕ → List (SubFace × CuTerm) → List (SubFace × CuTermNC) + + fillCongs fuel dim (hco x cu) = + hco (fillCongsS fuel x) (fillCongs fuel dim cu) + fillCongs fuel dim (cell x₁) = cell (liftVarsFrom 1 dim x₁) + fillCongs zero dim (𝒄ong' x x₁) = cell (R.lit (R.string "out of fuel in fillCongs")) + fillCongs (suc fuel) dim (𝒄ong' t []) = cell (liftVarsFrom 1 dim t) + fillCongs (suc fuel) dim (𝒄ong' t xs) = uncurry (congFill fuel dim) (t , xs) + + + fillCongsS fuel [] = [] + fillCongsS fuel ((sf , x) ∷ xs) = + (sf ∷ʳ nothing , fillCongs fuel (suc (sfDim sf)) x) ∷ fillCongsS fuel xs + + +open fillCongs using (fillCongs) public diff --git a/Cubical/Tactics/PathSolver/Degen.agda b/Cubical/Tactics/PathSolver/Degen.agda new file mode 100644 index 0000000000..7599fb9b1e --- /dev/null +++ b/Cubical/Tactics/PathSolver/Degen.agda @@ -0,0 +1,139 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.PathSolver.Degen where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Interpolate + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sigma + + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar + +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.QuoteCubical +open import Cubical.Tactics.Reflection.CuTerm +open import Cubical.Tactics.Reflection.Error + + +undegenTerm : Bool → ℕ → ℕ → R.Term → R.TC R.Term +undegenTerm onEnd offset dim = + atVarOrDefM.rv + (λ n k _ args → R.var (n + k) <$> args) + h + zero ∘S (if onEnd then (idfun _) else (liftVarsFrom 1 (offset + dim))) + + where + + g : R.Name → List (R.Arg R.Term) → R.Term → Maybe R.Term + g (quote _∨_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote _∧_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote ~_) a@(v[ _ ]) tm = just tm + g _ _ _ = nothing + + h : ℕ → + R.Name → + List (R.Arg R.Term) → R.TC (List (R.Arg R.Term)) → R.TC R.Term + h _ nm arg argM = + Mb.rec (R.def nm <$> argM) + ((extractIExprM >=& + (IExpr→Term + ∘ mapVarsInIExpr (_+ offset) + ∘ undegen onEnd dim + ∘ mapVarsInIExpr (_∸ offset) ))) + (g nm arg (R.def nm arg)) + +undegenTerm2 : Bool → ℕ → ℕ → R.Term → R.TC R.Term +undegenTerm2 onEnd offset dim = + atVarOrDefM.rv + (λ n k _ args → R.var (n + k) <$> args) + h + zero ∘S (if onEnd then (idfun _) else (liftVarsFrom 1 (offset + dim))) + + where + + g : R.Name → List (R.Arg R.Term) → R.Term → Maybe R.Term + g (quote _∨_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote _∧_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote ~_) a@(v[ _ ]) tm = just tm + g _ _ _ = nothing + + h : ℕ → + R.Name → + List (R.Arg R.Term) → R.TC (List (R.Arg R.Term)) → R.TC R.Term + h _ nm arg argM = + Mb.rec (R.def nm <$> argM) + ((extractIExprM >=& + (IExpr→Term + ∘ mapVarsInIExpr (_+ offset) + ∘ undegen2 onEnd dim + ∘ mapVarsInIExpr (_∸ offset) ))) + (g nm arg (R.def nm arg)) + + + +private + variable + ℓ : Level + A : Type ℓ + CongGuard : Type + + +module UndegenCell (dim : ℕ) where + + undegenCell : (R.Term × R.Term) → R.Term → R.TC R.Term + undegenCell (t0 , tI) t = do + let eai = (extractAllIExprs t0) + -- concatMapM (pure ∘S ("" ∷nl_) ∘S [_]ₑ ∘S IExpr→Term) eai >>= R.typeError + + fe ← undegenFcs dim eai --(extractAllIExprs t0) + + -- Mb.rec (pure t) + let ie = IExpr→Term (F→I dim fe) + pure + -- addNDimsToCtx (suc dim) $ R.typeError $ [_]ₑ $ + (R.def (quote hcomp) + (vlam "undegenCellDim" + (R.def (quote primPOr) + (R.def (quote ~_) v[ 𝒗 (suc dim) ] v∷ (R.def (quote _∨_) ((𝒗 (suc dim)) v∷ + v[ (liftVars ie) ])) v∷ + (constPartialR (R.def (quote ~_) v[ 𝒗 (suc dim) ]) (liftVarsFrom 1 (suc dim) tI)) + v∷ v[ constPartialR ((R.def (quote _∨_) ((𝒗 (suc dim)) v∷ + v[ (liftVars ie) ]))) (liftVars t) ])) v∷ v[ t ])) + where + constPartialR : R.Term → R.Term → R.Term + constPartialR tI tA = R.def (quote constPartial) (tA v∷ v[ tI ]) + + + mbUndegen : R.Term → R.TC (Maybe (R.Term × R.Term) × R.Term) + mbUndegen tm = do + + allNonDeg ← foldrM + (\ie b → (b and_) <$> (isNonDegen dim ie)) + true (extractAllIExprs tm) + if allNonDeg then (pure (nothing , tm)) else + do idt0 ← undegenTerm2 true zero dim tm + idt1 ← undegenTerm2 false zero dim tm + pure ( just (tm , idt1) , idt0) + + mbUndegen' : R.Term → R.TC (Maybe (R.Term × R.Term) × R.Term) + mbUndegen' tm = do + + allNonDeg ← foldrM + (\ie b → (b and_) <$> (isNonDegen dim ie)) + true (extractAllIExprs tm) + if allNonDeg then (pure (nothing , tm)) else + do idt0 ← undegenTerm true zero dim tm + idt1 ← undegenTerm false zero dim tm + pure ( just (tm , idt1) , idt0) + diff --git a/Cubical/Tactics/PathSolver/Macro.agda b/Cubical/Tactics/PathSolver/Macro.agda new file mode 100644 index 0000000000..1ae60bc8d3 --- /dev/null +++ b/Cubical/Tactics/PathSolver/Macro.agda @@ -0,0 +1,137 @@ +{- + +This module provides macros for manipulating cubical terms using reflection machinery. It includes utilities to transform terms by exposing functions from the `CongComp` module and simplify the construction of `hcomps`. + +Usage examples for these macros are available in the `Cubical.Tactics.PathSolver.MacroExamples` module. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Macro where + +open import Cubical.Foundations.Function + +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sigma + +open import Agda.Builtin.String + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar +import Agda.Builtin.Reflection as R + +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.Reflection.QuoteCubical +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Error +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.CuTerm + +open import Cubical.Tactics.Reflection.CuTerm public using (⁇) + +open import Cubical.Tactics.PathSolver.CongComp + +{- + +### Macros Exposing Helpers from `CongComp` Module + +These macros expose helpers from the `CongComp` module, offering broader application beyond the `NSolver` and `MonoidalSolver` macros. They support terms of arbitrary dimensions and complex interval expressions. + +- **Usage**: Apply these macros in a context where interval variables are present. Paths and n-cubes must be applied to the respective dimensions. +- **Functionality**: + - **`cong$` Macro**: Transforms the given term by bringing all `hcomps` to the top, so all functions are applied "all the way down" through the `hcomps`. This process happens iteratively for all subterms. + - **`cong$-fill` Macro**: Provides a filler path from the initial term to the term produced by `cong$`. Assumes there is an additional interval variable in the context not present in the processed term, used as the filling direction. + +-} + +macro + cong$ : R.Term → R.Term → R.TC Unit + cong$ t h = do + ty ← R.inferType t + t ← normaliseWithType ty t + cc ← getCuCtx + let dim = length cc + co ← R.getContext + cu ← R.inContext (drop dim co) $ appCongs dim + <$> quoteCuTerm (just (dropVars.rv dim zero ty)) dim (t) + let r = ToTerm.toTerm cc cu + R.unify r h <|> R.typeError [ r ]ₑ + + + +macro + cong$-fill : R.Term → R.Term → R.TC Unit + cong$-fill t h = do + ty ← R.inferType t + t ← normaliseWithType ty t + cc ← getCuCtx + let dim = predℕ (length cc) + co ← R.getContext + cu ← R.inContext (drop (suc dim) co) $ fillCongs 100 dim + <$> quoteCuTerm nothing --(just (dropVars.rv (suc dim) zero ty)) + dim (dropVar dim t) + let r = ToTerm.toTerm cc cu + + R.unify r h --<|> R.typeError [ r ]ₑ + + + +makeH? : R.Term → R.Term → R.TC String +makeH? iT eT = do + cuCtx ← getCuCtx + let dim = length cuCtx + fE ← iFxprFromSpec dim iT + + cuCtx ← L.map (map-snd (const nothing)) ∘S take dim <$> R.getContext + (((hcoFromIExpr dim fE eT) >>= codeGen.ppCT'' false dim cuCtx 10) >>= R.formatErrorParts) + + where + iFxprFromSpec : ℕ → R.Term → R.TC FExpr + iFxprFromSpec dim t = + (R.unquoteTC {A = ℕ} t <⊎> extractIExprM t) + <|> (R.typeError $ + "Wrong parameter!" + ∷nl [ "pass either ℕ or Interval Expr as first argument!" ]ₑ) + >>= pure ∘S ⊎.rec (allSubFacesOfSfDim dim ∘S min (predℕ dim)) + ((_++fe (allSubFacesOfSfDim dim 0)) ∘S I→F) + +{- + +The following macros are designed to facilitate the construction of homogeneous `hcomps` given a face expression. They simplify the process of writing nested `hcomp`'s . + +## Macros Overview + + **Usage**: These macros must be called in a context where interval variables are in scope. + **Arguments**: + **First Argument**: + - A face expression (e.g., `(i ∨ k) ∨ (~k ∧ ~j)`) or + - A natural number that specifies the dimension. This will generate a face expression consisting of all faces of the specified dimension from available interval variables. + **Second Argument**: + - A term to serve as a value for all generated cells, or + - The `⁇` symbol (`\??`) to insert holes into all cells for further editing. +-} + +macro + + + -- Produces the term directly in the hole where it is called (via `C-c C-m` in emacs). + -- Note: This macro may have issues correctly introducing holes in the face expressions. + h? : R.Term → R.Term → R.Term → R.TC Unit + h? iT eT h = do + src ← makeH? iT eT + unifyFromString src h + + + -- Generates a slightly better-looking term compared to `h?`. + -- The result is printed to the AgdaInfo buffer, + -- and users must manually copy it to the desired location in the code. + h?' : R.Term → R.Term → R.Term → R.TC Unit + h?' iT eT _ = do + src ← makeH? iT eT + R.typeError [ indent ' ' 8 src ]ₑ diff --git a/Cubical/Tactics/PathSolver/MacroExamples.agda b/Cubical/Tactics/PathSolver/MacroExamples.agda new file mode 100644 index 0000000000..438c73055b --- /dev/null +++ b/Cubical/Tactics/PathSolver/MacroExamples.agda @@ -0,0 +1,165 @@ +{- +This module provides usage examples for the macros defined in `Cubical.Tactics.PathSolver.Macro`. +Usage of macros is documented in `Cubical.Tactics.PathSolver.Macro` module. +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.MacroExamples where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Tactics.PathSolver.Path + +open import Cubical.Tactics.PathSolver.Macro +open import Cubical.Tactics.Reflection.QuoteCubical + + +private + variable + ℓ : Level + A B : Type ℓ + + + +module _ (SA : NPath 3 A) (f : A → B) where + open NPath SA + + f[assoc] : cong f 𝑝₀ ∙ cong f 𝑝₁ ∙ cong f 𝑝₂ + ≡ (cong f 𝑝₀ ∙ cong f 𝑝₁) ∙ cong f 𝑝₂ + f[assoc] i j = cong$ (f (assoc 𝑝₀ 𝑝₁ 𝑝₂ i j)) + + +module _ (SA : NPath 6 A) (f : A → {A} → A → A) (g : A → A) (𝑝ₓ : g (NPath.𝑣₀ SA) ≡ g (NPath.𝑣₀ SA)) where + open NPath SA + + p : f 𝑣₀ {g (𝑣₁)} 𝑣₁ ≡ f 𝑣₃ {g 𝑣₃} 𝑣₆ + p i = (f ((𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) i) {g ((𝑝₁ ∙' 𝑝₂) i)} ((𝑝₁ ∙∙ 𝑝₂ ∙∙ (𝑝₃ ∙∙ 𝑝₄ ∙∙ 𝑝₅)) i)) + + + _ : (λ i → cong$ (p i)) + ≡ + (λ i → f (𝑝₀ i) {g (𝑝₁ i)} (𝑝₁ i)) + ∙∙ (λ i → f (𝑝₁ i) {g (𝑝₂ i)} (𝑝₂ i)) + ∙∙ ((λ i → f 𝑣₂ {g 𝑣₃} (𝑝₃ i)) + ∙∙ (λ i → f (𝑝₂ i) {g 𝑣₃} (𝑝₄ i)) + ∙∙ λ i → f 𝑣₃ {g 𝑣₃} (𝑝₅ i)) + _ = refl + + cg² : ∀ {x y : A} → (x ≡ y) → g (g x) ≡ g (g y) + cg² = congS (g ∘S g) + + cpf : Square (cong g 𝑝₀) (cong g (𝑝₀ ∙ 𝑝₁)) + refl (cong g 𝑝₁) + cpf i j = g (compPath-filler 𝑝₀ 𝑝₁ i j) + + cpf' : Square (cong g 𝑝₀) (cong g 𝑝₀ ∙ cong g 𝑝₁) + refl (cong g 𝑝₁) + cpf' i j = cong$ (cpf i j) + + + cpf≡cpf'I : I → I → I → A + cpf≡cpf'I _ i j = cong$-fill (cpf i j) + + cpf≡cpf' : Cube + cpf cpf' + (λ i j → cpf≡cpf'I i i0 j) + (λ i j → cpf≡cpf'I i i1 j) + (λ i j → cpf≡cpf'I i j i0) + (λ i j → cpf≡cpf'I i j i1) + cpf≡cpf' _ i j = cong$-fill (cpf i j) + + + + cpf2 : Square (cong g (𝑝ₓ ∙ cong g (𝑝₀ ∙ 𝑝₁))) + (cong g ((𝑝ₓ ∙ cong g (𝑝₀ ∙ 𝑝₁)) ∙ cong g (𝑝₂ ∙ 𝑝₃))) + refl (cg² (𝑝₂ ∙ 𝑝₃)) + cpf2 i j = g (compPath-filler (𝑝ₓ ∙ cong g (𝑝₀ ∙ 𝑝₁)) (cong g (𝑝₂ ∙ 𝑝₃)) i j) + + cpf2' : Square + (cong g 𝑝ₓ ∙ cg² 𝑝₀ ∙ cg² 𝑝₁) + ((cong g 𝑝ₓ ∙ cg² 𝑝₀ ∙ cg² 𝑝₁) ∙ cg² 𝑝₂ ∙ cg² 𝑝₃) + refl + (cg² 𝑝₂ ∙ cg² 𝑝₃) + cpf2' i j = cong$ (cpf2 i j) + + + cpf2≡cpf2'I : I → I → I → A + cpf2≡cpf2'I _ i j = cong$-fill (cpf2 i j) + + cpf2≡cpf2' : Cube + cpf2 cpf2' + (λ i j → cpf2≡cpf2'I i i0 j) + (λ i j → cpf2≡cpf2'I i i1 j) + (λ i j → cpf2≡cpf2'I i j i0) + (λ i j → cpf2≡cpf2'I i j i1) + cpf2≡cpf2' z i j = cpf2≡cpf2'I z i j + + + +module _ (A : Type) (a : A) (p : a ≡ a) + (s : Square p p p p) where + + + {- + Examples below can be recreated by replacing the body of the definition with a hole, + then placing the example macro call in that hole and executing `C-c C-m` in Emacs. + (for h?' macro, result needs to be manually coppied from AgdaInfo buffer) + + results including holes are commented out to allow compilation of module + -} + + + -- -- h?' 1 ⁇ + -- c₀ : I → I → I → A + -- c₀ i j k = + -- hcomp (λ 𝒛₀ → λ { + -- (j = i0)(i = i0) → {!!} + -- ;(j = i0)(i = i1) → {!!} + -- ;(j = i1)(i = i0) → {!!} + -- ;(j = i1)(i = i1) → {!!} + -- ;(k = i0)(i = i0) → {!!} + -- ;(k = i0)(i = i1) → {!!} + -- ;(k = i0)(j = i0) → {!!} + -- ;(k = i0)(j = i1) → {!!} + -- ;(k = i1)(i = i0) → {!!} + -- ;(k = i1)(i = i1) → {!!} + -- ;(k = i1)(j = i0) → {!!} + -- ;(k = i1)(j = i1) → {!!} + -- }) + -- ({!!}) + + + -- h? 2 (s i (j ∧ k)) + c₁ : I → I → I → A + c₁ i j k = hcomp + (λ { 𝒛₀ (i = i0) → s i0 (j ∧ k) + ; 𝒛₀ (i = i1) → s i1 (j ∧ k) + ; 𝒛₀ (j = i0) → s i (i0 ∧ k) + ; 𝒛₀ (j = i1) → s i (i1 ∧ k) + ; 𝒛₀ (k = i0) → s i (j ∧ i0) + ; 𝒛₀ (k = i1) → s i (j ∧ i1) + }) + (s i (j ∧ k)) + + -- -- h? (i ∨ ~ j ∨ (~ i ∧ j)) ⁇ + -- c₂ : I → I → I → A + -- c₂ i j k = hcomp + -- (λ { 𝒛₀ (i = i1) → _ + -- ; 𝒛₀ (j = i0) → _ + -- ; 𝒛₀ (i = i0) (j = i1) → _ + -- }) + -- {!!} + + -- h?' (i ∨ ~ j ∨ (~ i ∧ j)) (s i (j ∧ k)) + c₃ : I → I → I → A + c₃ i j k = + hcomp (λ 𝒛₀ → λ { + (i = i1) → s i (j ∧ k) + ;(j = i0) → s i (j ∧ k) + ;(j = i1)(i = i0) → s i (j ∧ k) + }) + (s i (j ∧ k)) diff --git a/Cubical/Tactics/PathSolver/MonoidalSolver/Examples.agda b/Cubical/Tactics/PathSolver/MonoidalSolver/Examples.agda new file mode 100644 index 0000000000..e5c62b0539 --- /dev/null +++ b/Cubical/Tactics/PathSolver/MonoidalSolver/Examples.agda @@ -0,0 +1,201 @@ +{- + +This file provides concrete examples demonstrating the usage of the `solvePaths` and `simplifyPath` macros from the `MonoidalSolver` module. + +Examples here are crafted specificaly to showcase instances where monoidal-groupoid laws are +necessary (NSolver would fail on those). + +Higher dimensional examples are absent, +since this solver is capable only of handling squares at this point. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.MonoidalSolver.Examples where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + + +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sum + + +open import Agda.Builtin.String + + + +open import Cubical.Tactics.PathSolver.MonoidalSolver.MonoidalSolver +open import Cubical.Tactics.PathSolver.Path + + + +open import Cubical.Tactics.Reflection.CuTerm + + + +private + variable + ℓ ℓ' : Level + A B C D E F : Type ℓ + +module _ {ℓ} {A : Type ℓ} (SA : NPath 5 A) (f : A → A → A) + where + + +module E0' (SA : NPath 5 A) + (SB : NPath 5 B) + (SC : NPath 3 C) + (SD : NPath 3 D) where + + module A = NPath SA + module B = NPath SB + module C = NPath SC + module D = NPath SD + + + module _ (f f' : A → B → D) (h : A → A → A) + (p : f A.𝑣₂ B.𝑣₁ ≡ f' A.𝑣₃ B.𝑣₂) + (XX : A) + (q : A.𝑣₂ ≡ h A.𝑣₂ A.𝑣₂) where + + _ : (cong₂ f + (cong₂ h (A.𝑝₀ ∙ A.𝑝₁) (A.𝑝₀ ∙ A.𝑝₁) ∙ sym q) + (B.𝑝₀ ∙∙ B.𝑝₁ ∙∙ sym B.𝑝₁ ) + ∙∙ p + ∙∙ cong (f' A.𝑣₃) (B.𝑝₂ ∙ B.𝑝₃) + ∙∙ cong (flip f' B.𝑣₄) (A.𝑝₃ ∙ A.𝑝₄) + ∙∙ cong (f' A.𝑣₅) (sym B.𝑝₃) + ∙∙ cong (flip f' B.𝑣₃) (sym A.𝑝₄) + ∙∙ refl + ∙∙ cong (f' A.𝑣₄) (sym B.𝑝₂) + ∙∙ cong (flip f' B.𝑣₂) (sym A.𝑝₃) + ∙∙ sym p + ∙∙ cong (f A.𝑣₂) (sym B.𝑝₀) + ∙∙ cong (flip f B.𝑣₀) (q ∙ cong₂ h (sym A.𝑝₁) (sym A.𝑝₁)) + ∙∙ cong (flip f B.𝑣₀) (cong₂ h (sym A.𝑝₀) (sym A.𝑝₀))) + ≡ refl + _ = solvePaths + + module _ (f₄ : A → {B} → C → D → E) where + cong₄Funct∙₃ : + ([ f₄ ]$≡ A.𝑝₀ ≡$'≡ B.𝑝₀ ≡$≡ C.𝑝₀ ≡$≡ D.𝑝₀) + ∙ ([ f₄ ]$≡ A.𝑝₁ ≡$'≡ B.𝑝₁ ≡$≡ C.𝑝₁ ≡$≡ D.𝑝₁) + ∙ ([ f₄ ]$≡ A.𝑝₂ ≡$'≡ B.𝑝₂ ≡$≡ C.𝑝₂ ≡$≡ D.𝑝₂) + ≡ + ([ f₄ ]$≡ (A.𝑝₀ ∙ A.𝑝₁ ∙ A.𝑝₂) + ≡$'≡ (B.𝑝₀ ∙ B.𝑝₁ ∙ B.𝑝₂) + ≡$≡ (C.𝑝₀ ∙ C.𝑝₁ ∙ C.𝑝₂) + ≡$≡ (D.𝑝₀ ∙ D.𝑝₁ ∙ D.𝑝₂)) + cong₄Funct∙₃ = solvePaths + + cong-comm₄ : + (λ i → f₄ (A.𝑝₀ i) {B.𝑣₀} (C.𝑣₀) (D.𝑣₀)) + ∙ (λ i → f₄ (A.𝑣₁) {B.𝑝₀ i} (C.𝑣₀) (D.𝑣₀)) + ∙ (λ i → f₄ (A.𝑣₁) {B.𝑣₁} (C.𝑝₀ i) (D.𝑣₀)) + ∙ (λ i → f₄ (A.𝑣₁) {B.𝑣₁} (C.𝑣₁) (D.𝑝₀ i)) + ≡ + (λ i → f₄ (A.𝑣₀) {B.𝑣₀} (C.𝑣₀) (D.𝑝₀ i)) + ∙ (λ i → f₄ (A.𝑣₀) {B.𝑣₀} (C.𝑝₀ i) (D.𝑣₁)) + ∙ (λ i → f₄ (A.𝑣₀) {B.𝑝₀ i} (C.𝑣₁) (D.𝑣₁)) + ∙ (λ i → f₄ (A.𝑝₀ i) {B.𝑣₁} (C.𝑣₁) (D.𝑣₁)) + cong-comm₄ = solvePaths + + + + _ : (f : A → B → C) → Square + (cong₂ f A.𝑝₀ B.𝑝₁) + (cong₂ f (sym A.𝑝₂) (B.𝑝₀ ∙ B.𝑝₁ ∙ B.𝑝₂)) + (cong₂ f A.𝑝₀ (sym B.𝑝₀) ∙ cong₂ f A.𝑝₁ B.𝑝₀ ∙ cong₂ f A.𝑝₂ (sym B.𝑝₀)) + (cong₂ f A.𝑝₁ B.𝑝₂) + _ = λ f → solvePaths + + + +module simplify-examples {x y z w : A} {x' y' z' w' : B} + (p : x ≡ y) + (q : y ≡ z) + (r : z ≡ w) + (p' : x' ≡ y') + (q' : y' ≡ z') + (r' : z' ≡ w') + (f : A → A) (f₂ : A → A → A) (f₄ : A → {A} → A → A → A) where + + + e0 : _ ≡ (p ∙' (q ∙' r)) + e0 = simplifyPath ((p ∙∙ q ∙∙ sym q ) ∙∙ q ∙∙ r) + + + e1 : _ ≡ p ∙' (q ∙' r) + e1 = simplifyPath (p ∙∙ q ∙∙ r ) + + e1' : _ ≡ q ∙' r + e1' = simplifyPath (refl ∙∙ q ∙∙ r ) + + + e2 : _ ≡ p + e2 = simplifyPath (p ∙∙ refl ∙∙ refl ) + + + + e3 : _ ≡ ((λ 𝓲 → f (p 𝓲)) ∙' ((λ 𝓲 → f (q 𝓲)) ∙' (λ 𝓲 → f (r 𝓲)))) + e3 = simplifyPath (cong f p ∙ cong f q ∙ (refl ∙ cong f r)) + + e4 : _ ≡ cong₂ f₂ q p + e4 = simplifyPath (cong (f₂ y) p ∙ cong (flip f₂ y) q ) + + + + e5 : _ ≡ λ 𝓲 → f₄ (p 𝓲) {q 𝓲} (r 𝓲) (q 𝓲) + e5 = simplifyPath + ((λ i → f₄ (p i) {y} z (p (~ i))) + ∙∙ (λ i → f₄ y {q i} z ((p ∙ q) i)) ∙∙ + (λ i → f₄ ((refl {x = y} ∙' refl {x = y}) i) {z} (r i) z) ) + + + + +module _ (A B C D : Type ℓ) + (A≃B : A ≃ B) (B≃C : B ≃ C) (C≃D : C ≃ D) + where + + _ : Square (cong List (cong₂ _⊎_ (ua A≃B) (sym (ua B≃C))) ∙ cong List (cong (_⊎ B) (ua B≃C))) + (cong List (cong₂ _⊎_ (ua A≃B ∙∙ ua B≃C ∙∙ ua C≃D) (sym (ua B≃C)))) + refl (cong List (cong (_⊎ B) (ua C≃D))) + _ = solvePaths + + + + module _ where + open import Cubical.Data.Prod + + -- if we switch to Σ based definition of product, the following will fail due + -- to the problem with handling constant lambdas in the solver (TODO) + -- open import Cubical.Data.Sigma + + + _ : Square (cong List (cong₂ _×_ (ua A≃B) (sym (ua B≃C))) ∙ cong List (cong (_× B) (ua B≃C))) + (cong List (cong₂ _×_ (ua A≃B ∙∙ ua B≃C ∙∙ ua C≃D) (sym (ua B≃C)))) + refl (cong List (cong (_× B) (ua C≃D))) + _ = solvePaths + + + + +-- proving this requires little bit of gymnastics: +-- 1. using _$sp_ wich is version of _$_ wich is prevented from normalisation inside the macro +-- 2. suplying H as equality of functions + +homotopyNatural' : {B : Type ℓ'} {f g : A → B} (H : ∀ a → f a ≡ g a) {x y : A} (p : x ≡ y) → + H x ∙ cong g p ≡ cong f p ∙ H y +homotopyNatural' {f = f} {g = g} H {x} {y} p = h (funExt H) + where + h : (H' : f ≡ g) → (cong (_$sp x) H') ∙ cong (g $sp_) p ≡ cong (f $sp_) p ∙ (cong (_$sp y) H') + h H' = solvePaths + diff --git a/Cubical/Tactics/PathSolver/MonoidalSolver/MonoidalSolver.agda b/Cubical/Tactics/PathSolver/MonoidalSolver/MonoidalSolver.agda new file mode 100644 index 0000000000..7a92ff3b6a --- /dev/null +++ b/Cubical/Tactics/PathSolver/MonoidalSolver/MonoidalSolver.agda @@ -0,0 +1,218 @@ +{- + +This module implements a solver capable of handling squares, which are boundaries built of one-dimensional faces. At this stage, it does not support higher-dimensional cubes. + +**Simplified One-Dimensional View**: + - Represents compositions involving only one interval variable. + - Defined using the specialized `1DimView` type. + - Equipped with instances of both `RawMonad` and `RawApplicative`. + - Simplifies the manipulation and normalization of paths within the solver + (compared to more general, arbitrary dimensional CuTerm representation) + +### Solver steps: + +1. **Identifying Boundaries**: + - The solver begins by identifying the boundary of a square term. + +2. **Quoting into Specialized Representation**: + - The boundary terms are quoted into a specialized one-dimensional representation using the `quote1D` function. + +3. **Applying Generalized `cong-∙`**: + - (if necessary) Uses the `fillCongs` functions from the `CongComp` module to "push to the bottom" all application of functions (bring all the composiitons to the "top") + +4. **Filler Construction**: + - Constructs fillers for each face using the `_↙_` operator from the `PathEval` module. + - Relies on the aspects of the free monoidal groupoid structure as implemented in `PathEval`. + +5. **Assembling the Final Square**: + - The final square term is assembled using the `◪mkSq` lemma from `Path.agda`. + +-} + +{-# OPTIONS --safe #-} +module Cubical.Tactics.PathSolver.MonoidalSolver.MonoidalSolver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat + + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar +import Agda.Builtin.Reflection as R + +open import Cubical.Tactics.Reflection + +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Variables + +open import Cubical.Tactics.PathSolver.CongComp + +open import Cubical.Tactics.Reflection.QuoteCubical +open import Cubical.Tactics.Reflection.Error +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.CuTerm + +open import Cubical.Tactics.PathSolver.Degen +open import Cubical.Tactics.PathSolver.MonoidalSolver.PathEval +open import Cubical.Tactics.PathSolver.Path + +private + variable + ℓ : Level + A B : Type ℓ + + + + + +flipOnFalse : Bool → R.Term → R.Term +flipOnFalse b t = if b then t else R.def (quote ~_) v[ t ] + + + +data 1DimView (A : Type ℓ) : Type ℓ where + [_]1d : A → 1DimView A + [_∙∙1d_∙∙1d_]1d : 1DimView A → 1DimView A → 1DimView A → 1DimView A + + + + +map-1DimView : (f : A → B) → 1DimView A → 1DimView B +map-1DimView f [ x ]1d = [ f x ]1d +map-1DimView f [ x ∙∙1d x₁ ∙∙1d x₂ ]1d = [ (map-1DimView f x) ∙∙1d (map-1DimView f x₁) ∙∙1d (map-1DimView f x₂) ]1d + + + +module _ {M : Functorω} {{_ : RawApplicative M}} {{_ : RawMonad M}} where + + mapM-1DimView : (f : A → M B) → 1DimView A → M (1DimView B) + mapM-1DimView f [ x ]1d = ⦇ [ f x ]1d ⦈ + mapM-1DimView f [ x ∙∙1d x₁ ∙∙1d x₂ ]1d = + ⦇ [ (mapM-1DimView f x) ∙∙1d (mapM-1DimView f x₁) ∙∙1d (mapM-1DimView f x₂) ]1d ⦈ + + +sym1DimView : 1DimView (A × R.Term) → 1DimView (A × R.Term) +sym1DimView [ (a , x) ]1d = [ (a , (invVar zero x)) ]1d +sym1DimView [ x ∙∙1d x₁ ∙∙1d x₂ ]1d = + [ (sym1DimView x₂) ∙∙1d (sym1DimView x₁) ∙∙1d (sym1DimView x) ]1d + +module _ (a : A) where + to1DimView : CuTerm' ⊥ A → Maybe (1DimView (A × R.Term)) + + to1DimView (hco (((just b) ∷ [] , fc) ∷ ((just b') ∷ [] , fc') ∷ []) x₁) = + (do x₁ ← to1DimView x₁ + f ← to1DimView fc + f' ← to1DimView fc' + let f0 = if b then f' else f + let f1 = if b then f else f' + just [ sym1DimView f0 ∙∙1d x₁ ∙∙1d f1 ]1d) + + + to1DimView (cell' x x₁) = just [ (x , x₁) ]1d + to1DimView _ = nothing + + +1dvEnd : 1DimView (A × R.Term) → R.TC PathTerm +1dvEnd y@([ (_ , x) ]1d) = 𝒓efl <$> (addNDimsToCtx 1 $ R.normalise + (replaceVarWithCon (λ { zero → just (quote i1) ; _ → nothing }) x)) +1dvEnd y@([ x ∙∙1d x₁ ∙∙1d x₂ ]1d) = 1dvEnd x₂ + + + +fill1DV' : 1DimView (Maybe (R.Term × R.Term) × PathTerm) → PathTerm → R.TC (R.Term × PathTerm) +fill1DV' [ nothing , p ]1d q = do + (q , squareTerm s) ← p ↙ q + pure (s , q) +fill1DV' [ just (_ , ud≡) , p ]1d q = do + (q , squareTerm s) ← p ↙ q + pure (R.def (quote comp₋₀) (s v∷ v[ vlam "𝓳" $ vlam "𝓲" ud≡ ]) , q) +fill1DV' [ p₀ ∙∙1d p₁ ∙∙1d p₂ ]1d q = do + (s₂ , q) ← fill1DV' p₂ q + (s₁ , q) ← fill1DV' p₁ q + (s₀ , q) ← fill1DV' p₀ q + pure (R.def (quote _∙∙■_∙∙■_) (s₀ v∷ s₁ v∷ v[ s₂ ]) , q ) + +fill1DV : 1DimView (A × R.Term) → R.TC (R.Term × PathTerm) +fill1DV x = + + 1dvEnd x >>= (fill1DV' (map-1DimView (λ (_ , pt) → nothing , asPathTerm pt) x)) + + + +quote1D : Maybe R.Type → R.Term → R.TC ((Maybe R.Term) × 1DimView (Maybe (R.Term × R.Term) × R.Term)) +quote1D mbty t = do + cu ← extractCuTermFromPath mbty t + let cu' = appCongs 1 cu + congPa ← pure (ToTerm.toTerm (defaultCtx 2) (fillCongs 100 1 cu)) + let mbCongPa = if (hasVar 1 congPa) then just congPa else nothing + Mb.rec (R.typeError [ "imposible in simplifyPath" ]ₑ) + (λ (x , y) → x ,_ <$> mapM-1DimView (UndegenCell.mbUndegen' 1 ∘S snd) y) + (map-Maybe (mbCongPa ,_) (to1DimView _ cu')) + + + +simplifyFillTerm : Maybe R.Type → R.Term → R.TC R.Term +simplifyFillTerm mbTy t = do + (_ , 1dv) ← quote1D mbTy t + (s , _) ← fill1DV 1dv + pure s + -- (R.typeError $ [ s ]ₑ) + + + +macro + simplifyFill : R.Term → R.Term → R.TC Unit + simplifyFill t h = do + sq ← simplifyFillTerm nothing t + R.unify sq h + + simplifyPath : R.Term → R.Term → R.TC Unit + simplifyPath t h = do + sq ← simplifyFillTerm nothing t + R.unify (R.def (quote ◪→≡) v[ sq ]) h + +stepSq : R.Type → R.Term → Maybe PathTerm → R.TC (R.Term × PathTerm) +stepSq A p mbQ = do + (mbCong≡ , 1dt) ← quote1D (just A) $ vlam "𝒾" p + + q ← Mb.rec (1dvEnd 1dt) pure mbQ + (s , q') ← fill1DV' (map-1DimView (map-snd asPathTerm) 1dt ) q + + let s' = Mb.rec s + (λ c≡ → R.def (quote comp₋₀) (s v∷ v[ vlam "𝓳" $ vlam "𝓲" c≡ ])) + mbCong≡ + pure $ s' , q' + + +_$sp_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → A → B +f $sp a = f a + +macro + + + solvePaths : R.Term → R.TC Unit + solvePaths h = R.withReduceDefs (false , quote _$sp_ ∷ [ quote ua ]) do + hTy ← R.inferType h >>= wait-for-term >>= R.normalise + bdTM@(A , ((a₀₋ , a₁₋) , (a₋₀ , a₋₁))) ← (matchSquare <$> matchNCube hTy) >>= + Mb.rec (R.typeError [ "not a square" ]ₑ) pure + (a₁₋' , p₁₀) ← stepSq A a₁₋ nothing + (a₋₁' , p₀₁) ← stepSq A a₋₁ nothing + (a₀₋' , _) ← stepSq A a₀₋ (just p₀₁) + (a₋₀' , _) ← stepSq A a₋₀ (just p₁₀) + + let solution = R.def (quote ◪mkSq) + (a₀₋' v∷ a₁₋' v∷ a₋₀' v∷ a₋₁' v∷ []) + + R.unify solution h <|> R.typeError [ solution ]ₑ + diff --git a/Cubical/Tactics/PathSolver/MonoidalSolver/PathEval.agda b/Cubical/Tactics/PathSolver/MonoidalSolver/PathEval.agda new file mode 100644 index 0000000000..6e3c81dcbe --- /dev/null +++ b/Cubical/Tactics/PathSolver/MonoidalSolver/PathEval.agda @@ -0,0 +1,406 @@ +{- + +This module provides an implementation that can be perceived as a free monoidal groupoid generated by terms. It works under the assumption that there is only one interval variable in the context and that terms are free of composition operations. The core functionality involves decomposing terms and traversing them to identify subterms containing interval variables. These subterms are then marked as "paths" by wrapping them into temporary definitions. + +### Steps + +- **Path Identification**: Identifies subterms containing interval variables and marks them as "paths." +- **Term Wrapping**: Wraps identified "path" subterms into temporary definitions (`PathWrap`, `FillWrap`, `CompWrap`). +- **Normalization**: Treats terms with multiple subterms identified as "paths" as functors applied to the product of these subterms. Terms prepared in that way can be normalized according to monoidal-groupoid laws, generalizing the `cong₂Funct` lemma from the `Foundations.GroupoidLaws` module. + + +### Specialized Definitions Used only temporairly as markers during processing AST + +- **`ASTMarkers`**: Defines markers for `PathWrap`, `FillWrap`, and `CompWrap`. +- **Term Patterns**: Various term patterns (`fw`, `pw`, `cwd`, etc.) are used for identifying and manipulating wrapped terms. + + +### Utilities** +Functions like `intervalTest`, `transpose`, and `reduceComps` are utility functions aiding the main functionalities. + +At the end, whole process is wrapped into _↙_ operator, which can be seen as analogue of `compPath-filler` from prelude, but operating only on already normalised paths, and with composition operation resulting in "normalised" path according to monoidal-groipoid laws. + +-} + +{-# OPTIONS --safe -v testMarkVert:3 -v tactic:3 #-} + +module Cubical.Tactics.PathSolver.MonoidalSolver.PathEval where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ + +open import Cubical.Data.Sigma.Base + + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) + +import Agda.Builtin.Reflection as R + +open import Cubical.Tactics.PathSolver.Reflection +open import Cubical.Tactics.PathSolver.Path + +open import Agda.Builtin.String + +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Error +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.QuoteCubical + + +data ASTMarkers : Type where + PathWrap' FillWrap' CompWrap' : ASTMarkers + +PathWrap FillWrap CompWrap : ASTMarkers +PathWrap = PathWrap' +FillWrap = FillWrap' +CompWrap = CompWrap' + +WTerm = R.Term +CTerm = R.Term + + +pattern fw[1,_] x = R.def (quote FillWrap) (R.lit (R.name (quote true)) v∷ v[ x ]) +pattern fw[0,_] x = R.def (quote FillWrap) (R.lit (R.name (quote false)) v∷ v[ x ]) + + +pattern pw[_] x = R.def (quote PathWrap) (x v∷ []) +pattern pwd args = R.def (quote PathWrap) args + + +pattern cwd args = R.def (quote CompWrap) args + +pattern cw[] = R.def (quote CompWrap) [] +pattern cw[_] x = R.def (quote CompWrap) (x v∷ []) +pattern cw xs = R.def (quote CompWrap) xs +pattern _∷cw_ x xs = R.def (quote CompWrap) (x v∷ xs) + + +intervalTest : ℕ → R.Term → Bool +intervalTest _ (R.def (quote _∨_) _) = true +intervalTest _ (R.def (quote _∧_) _) = true +intervalTest _ (R.def (quote ~_) _) = true +intervalTest n (R.var k []) = n =ℕ k +intervalTest _ _ = false + + +wrapPaths : R.Term → WTerm +wrapPaths = atVarOrConM' f h g + where + f : ℕ → ℕ → List (R.Arg R.Term) → Maybe R.Term + f n v args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just pw[ (R.var (v + n) args) ] + else nothing + + h : ℕ → R.Name → List (R.Arg R.Term) → Maybe R.Term + h n nm args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just pw[ (R.con nm args) ] + else nothing + + g : ℕ → R.Name → List (R.Arg R.Term) → Maybe R.Term + g n nm args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just pw[ (R.def nm args) ] + else nothing + +wrapFills : R.Term → WTerm +wrapFills = atVarOrConM' f h g + where + f : ℕ → ℕ → List (R.Arg R.Term) → Maybe R.Term + f n v args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just fw[1, pw[ (R.var (v + n) args) ] ] + else nothing + + h : ℕ → R.Name → List (R.Arg R.Term) → Maybe R.Term + h n nm args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just fw[1, pw[ (R.con nm args) ] ] + else nothing + + + g : ℕ → R.Name → List (R.Arg R.Term) → Maybe R.Term + g n nm args = + if any? (L.map (intervalTest n ∘S unArg) args) + then just fw[1, pw[ (R.def nm args) ] ] + else nothing + + + +dropPathWraps : CTerm → R.Term +dropPathWraps = atVarOrDefM {{_}} {{RawMonadIdentityM}} + (λ _ v _ argsM → R.var v argsM) + λ _ d _ argsM → w (R.def d argsM) + + where + w : R.Term → R.Term + w pw[ x ] = x + w x = x + + +absorb : ℕ → WTerm → CTerm → R.TC CTerm + + +absorbStep : ℕ → WTerm → WTerm → R.TC (Maybe CTerm) +absorbStep n (cwd _) _ = R.typeError [ "cwd in absorbStep" ]ₑ +absorbStep n _ (cwd _) = R.typeError [ "cwd in absorbStep" ]ₑ +absorbStep zero pw[ x ] pw[ y ] = do + -- R.debugPrint "testMarkVert" 3 $ "-----" ∷nl x ∷nl "** \n" ∷nl [ y ]ₑ + (if_then (just fw[0, y ]) else nothing) <$> unifyTest (suc zero) x (invVar zero y) +absorbStep (suc _) pw[ x ] pw[ y ] = + R.typeError [ "absorbStep: todo - paths under abstraction" ]ₑ +absorbStep n x pw[ y ] = pure nothing +absorbStep n pw[ x ] y = pure nothing +absorbStep n (pwd _) _ = R.typeError [ "pwd1 in absorbStep!" ]ₑ +absorbStep n _ (pwd _) = R.typeError [ "pwd2 in absorbStep!" ]ₑ +absorbStep n x y = just <$> h x y + where + + hs : R.Sort → R.Sort → R.TC R.Sort + h : WTerm → WTerm → R.TC CTerm + + ha : List (R.Arg R.Term) → List (R.Arg R.Term) → R.TC (List (R.Arg R.Term)) + ha [] [] = pure [] + ha (R.arg ax x ∷ xs) (R.arg _ x' ∷ xs') = + ⦇ ⦇ R.arg ⦇ ax ⦈ (absorb n x x') ⦈ ∷ ha xs xs' ⦈ + ha _ _ = R.typeError [ "absorbStep: ha-failed" ]ₑ + + h (R.var x args) (R.var _ args') = R.var x <$> ha args args' + h (R.con c args) (R.con _ args') = R.con c <$> ha args args' + h (R.def f args) (R.def _ args') = R.def f <$> ha args args' + h (R.lam v (R.abs ai t)) (R.lam v' (R.abs _ t')) = + ⦇ R.lam ⦇ v ⦈ ⦇ R.abs ⦇ ai ⦈ (absorb (suc n) t t') ⦈ ⦈ + h (R.pat-lam cs args) (R.pat-lam cs' args') = R.typeError [ "absorbStep: todo - patLamb" ]ₑ + h (R.pi (R.arg ai a) (R.abs bi b)) (R.pi (R.arg ai' a') (R.abs bi' b')) = + ⦇ R.pi ⦇ R.arg ⦇ ai ⦈ (absorb n a a') ⦈ ⦇ R.abs ⦇ bi ⦈ (absorb (suc n) b b') ⦈ ⦈ + h (R.agda-sort s) (R.agda-sort s') = R.agda-sort <$> hs s s' + h (R.lit l) (R.lit l') = pure (R.lit l) + h (R.meta x x₂) (R.meta x' x₂') = R.typeError [ "absorbStep: todo - meta" ]ₑ + h R.unknown R.unknown = ⦇ R.unknown ⦈ + h t t' = R.typeError + $ "absorbStep: h-failed" ∷nl t ∷nl "---" ∷nl [ t' ]ₑ + + hs (R.set t) (R.set t') = R.set <$> absorb n t t' + hs (R.lit n) (R.lit _) = pure (R.lit n) + hs (R.prop t) (R.prop t') = R.prop <$> absorb n t t' + hs (R.propLit n) (R.propLit n') = pure (R.propLit n) + hs (R.inf n) (R.inf n') = pure (R.inf n) + hs R.unknown R.unknown = pure (R.unknown) + hs _ _ = R.typeError [ "absorbStep: hs-failed" ]ₑ + +absorbStep' : ℕ → WTerm → WTerm → R.TC (Maybe CTerm) +absorbStep' n x y = + w (hasVar zero x) (hasVar zero y) + + where + w : Bool → Bool → R.TC (Maybe CTerm) + w true true = absorbStep n x y + w true false = pure $ just (wrapFills (dropPathWraps x)) + w false false = pure (just x) + w false true = pure (just y) + +absorb _ _ cw[] = R.typeError [ "cw[] in absorb!" ]ₑ +absorb _ _ (cw[ y ]) = R.typeError [ "cw[_] in absorb!" ]ₑ +absorb n x (y ∷cw ys) = + absorbStep' n x y >>= + Mb.rec (pure (cw (fw[1, x ] v∷ y v∷ ys))) (pure ∘S _∷cw ys) +absorb n x y = absorbStep' n x y >>= + Mb.rec (pure (fw[1, x ] ∷cw v[ y ])) + pure + + +cTermEnd : CTerm → R.TC (Maybe WTerm) +cTermEnd = fixMb ∘S + atVarOrDefM (λ _ v _ argM → R.var v <$> argM) + (λ n def _ argsM → ((R.def def) <$> argsM) >>= reduceComps n) ∘S evFills + + where + evFills : CTerm → CTerm + evFills = + atVarOrM + (λ _ _ _ → nothing ) + λ _ nm args → atD (R.def nm args) + + where + atD : R.Term → Maybe R.Term + atD fw[1, x ] = just x + atD fw[0, x ] = just (replaceVarWithCon (λ { zero → just (quote i1) ; _ → nothing }) x) + atD _ = nothing + + reduceComps : ℕ → R.Term → R.TC R.Term + reduceComps _ cw[] = R.typeError [ "cTermEnd : reduceComps : unexpected cw[]" ]ₑ + reduceComps n cw[ p ] = pure p --if (hasVar n p) then else {!!} + reduceComps n t@(p ∷cw ps@((R.arg _ ps0) ∷ _)) = + if (hasVar n p) then pure t else + (pure (if length ps =ℕ 1 then ps0 else cw ps)) + reduceComps _ x = pure x + + fixMb : R.TC WTerm → R.TC (Maybe WTerm) + fixMb x = x >>= λ x → pure $ if (hasVar 0 x) then just x else nothing + +data FillWrapEval : Type where + headFW dropFW : FillWrapEval + +dropFillWraps : FillWrapEval -> CTerm → R.Term +dropFillWraps fwe = atVarOrDefM {{_}} {{RawMonadIdentityM}} + (λ _ v _ argsM → R.var v argsM) + λ _ d _ argsM → w fwe (R.def d argsM) + + where + + lift0Dim = remapVars λ { zero → suc zero ; n → n } + + w : FillWrapEval → R.Term → R.Term + -- w offsetFW fw[1, x ] = lift0Dim x + -- w offsetFW fw[0, x ] = invVar 1 (lift0Dim x) + w headFW fw[1, x ] = replaceVarWithTerm + (λ { zero → just (R.def (quote _∨_) (𝒗 1 v∷ v[ 𝒗 0 ])) + ; _ → nothing }) x + w headFW fw[0, x ] = replaceVarWithTerm + (λ { zero → just (R.def (quote _∨_) ((R.def (quote ~_) v[ 𝒗 1 ]) v∷ v[ 𝒗 0 ])) ; _ → nothing }) x + w dropFW fw[1, x ] = x + w dropFW fw[0, x ] = x + w _ x = x + + +transpose : ∀ {ℓ} {A : Type ℓ} → A → List (List A) → List (List A) +transpose default [] = [ [] ] +transpose default xss@(xs ∷ _) = + L.map (λ (k , ys) → L.map (λ ys → lookupAlways default ys k) xss ) (zipWithIndex xs) + +fill-flatten' : CTerm → List R.Term +fill-flatten' = hTop ∘S atVarOrConOrDefMmp + (λ _ v _ _ args' → R.var v <$> (h args')) + (λ _ c _ _ args' → R.con c <$> (h args')) + df ∘S dropPathWraps ∘S liftVarsFrom 1 1 + + where + + + + fill-offsetPa' : ℕ → List (R.Arg R.Term) → List (R.Arg R.Term) + fill-offsetPa' n xs = + let hd = fromJust-def (varg (R.lit (R.string "fatal in PathEval - offsetPa'"))) + (lookupMb xs zero) + hs* = mapArg (dropFillWraps headFW) hd + hd' = mapArg + (replaceVarWithCon (λ { zero → just (quote i0) ; _ → nothing })) hs* + in repeat (n ∸ length xs ) hd' ++ + hs* ∷ L.map (mapArg (dropFillWraps dropFW)) (tail xs) + + + h : List (List (R.Arg R.Term)) → List (List (R.Arg R.Term)) + h xs = + let maxL = foldr (max ∘S length) 1 xs + xs' = L.map (fill-offsetPa' maxL) xs + in transpose ((varg (R.lit (R.string "fatal in PathEval - flatten")))) xs' + + hTop : List R.Term → List R.Term + hTop = L.map (Mb.fromJust-def ( (R.lit (R.string "imposible in fill-flatten'")) ) + ∘S map-Maybe (unArg) ∘S flip lookupMb zero) ∘S h ∘S [_] ∘S L.map varg + + df : ℕ → + R.Name → + List (R.Arg R.Term) → + List (List (R.Arg R.Term)) → + List (List (R.Arg R.Term)) → List R.Term + df _ (quote CompWrap) _ _ args' = unArg <$> join args' + df _ nm _ _ args' = R.def nm <$> (h args') + + + +foldPath : List R.Term → R.Term +foldPath [] = q[ refl ] +foldPath (x ∷ []) = vlam "𝓲" x +foldPath (x ∷ xs@(_ ∷ _)) = R∙' (vlam "𝓲" x) (foldPath xs) + +foldPath' : List R.Term → Maybe R.Term +foldPath' [] = nothing +foldPath' (x ∷ []) = just $ vlam "𝓲" x +foldPath' (x ∷ xs@(_ ∷ _)) = just $ R∙' (vlam "𝓲" x) (foldPath xs) + +fillHeadTrm : R.Term → Maybe R.Term → R.TC R.Term +fillHeadTrm p nothing = pure (vlam "𝒋" (vlam "𝒊" p)) +fillHeadTrm p (just q) = do + p₀ ← hasVar 0 + <$> (addNDimsToCtx 2 $ R.normalise + (replaceVarWithCon (λ { (suc zero) → just (quote i0) ; _ → nothing }) p)) + p₁ ← hasVar 0 <$> (addNDimsToCtx 2 $ R.normalise + (replaceVarWithCon (λ { (suc zero) → just (quote i1) ; _ → nothing }) p)) + h p₀ p₁ + + where + h : Bool → Bool → R.TC R.Term + h false false = R.typeError [ "imposible in fillHeadTrm" ]ₑ + h false true = pure $ R.def (quote _∙f1_) (vlam "𝒋" (vlam "𝒊" p) v∷ v[ vlam "𝒋" q ]) + h true false = pure $ R.def (quote _∙f0_) (vlam "𝒋" (vlam "𝒊" p) v∷ v[ vlam "𝒋" q ]) + h true true = pure $ vlam "𝒋" (R∙' (vlam "𝓲" p) q) + + + +PathTerm = R.Term ⊎ R.Term + +pattern 𝒓efl x = inl x +pattern 𝒑λ x = inr x + + +record SquareTerm : Type where + constructor squareTerm + field + term : R.Term + + +asPathTerm : R.Term → PathTerm +asPathTerm tm = + if (hasVar zero tm) then (𝒑λ tm) else (𝒓efl tm) + + +bfs' : CTerm → R.TC R.Term +bfs' xs = do + let q = (foldPath' (tail (fill-flatten' xs))) + hd ← Mb.rec (R.typeError [ "imposible tfct≡" ]ₑ ) + pure (listToMaybe (fill-flatten' xs)) + -- addNDimsToCtx 2 $ R.typeError [ hd ]ₑ + fillHeadTrm hd q + + + + +-- compPath'-filler, but composition is 'simplified' according to groupoid laws +-- (p : x ≡ y) → (q : y ≡ z) → (Σ (p∙q ∈ x ≡ z) (Square q p∙q p refl)) +-- assumes that terms are already pre processed ussing : addNDimsToCtx 1 ∘S R.normalise ∘S pathApp + +_↙_ : PathTerm → PathTerm → R.TC (PathTerm × SquareTerm) +𝒓efl x ↙ q = q ,_ <$> (squareTerm <$> bfs' (⊎.rec (idfun _) (idfun _) q)) +𝒑λ x ↙ 𝒓efl y = + (𝒑λ (wrapPaths x) ,_) <$> (squareTerm <$> (bfs' (wrapFills x)) ) +𝒑λ p ↙ 𝒑λ q = do + pq-sq ← (absorb 0 (wrapPaths p) q) + + pq ← (cTermEnd pq-sq) >>= Mb.rec + ( 𝒓efl <$> (addNDimsToCtx 1 $ R.normalise + (replaceVarWithCon (λ { zero → just (quote i0) ; _ → nothing }) p))) (pure ∘S 𝒑λ) + + pq ,_ <$> (squareTerm <$> bfs' pq-sq) + + +macro + ↙-test : R.Term → R.Term → R.Term → R.TC Unit + ↙-test p q h = do + p' ← asPathTerm <$> (addNDimsToCtx 1 ∘S R.normalise ∘S pathApp) p + q' ← asPathTerm ∘S wrapPaths <$> (addNDimsToCtx 1 ∘S R.normalise ∘S pathApp) q + pq ← (SquareTerm.term ∘S snd) <$> p' ↙ q' + R.unify pq h diff --git a/Cubical/Tactics/PathSolver/NSolver/Examples.agda b/Cubical/Tactics/PathSolver/NSolver/Examples.agda new file mode 100644 index 0000000000..764692a01a --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Examples.agda @@ -0,0 +1,292 @@ +{- + +This file serves as a demonstration of the NSolver module, showcasing practical examples of its usage. It provides concrete cases for computing fillings of n-dimensional cubes under the assumption that boundary cells are expressible as paths (potentialy applied to arbitrary complex interval expressions). + +Unlike the `Tests` folder, which focuses on verifying correctness through unit tests, this file emphasizes illustration and practical application. + +Limitations of this solver, are best exemplified in the `Cubical.Tactics.PathSolver.MonoidalSolver.Examples`, +containing examples of complementary solver usage, capable of applying monoidal groupoid laws. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.NSolver.Examples where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sum + +open import Cubical.Data.Sigma + +open import Cubical.Tactics.PathSolver.NSolver.NSolver +open import Cubical.Tactics.PathSolver.Path + + +private + variable + ℓ : Level + A B : Type ℓ + + + +module Coherence (SA : NPath 7 A) where + open NPath SA + + + + a₀₋₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) 𝑝₀ (𝑝₂ ∙ 𝑝₃) + a₀₋₋ = solvePaths + + a₁₋₋ : Square (𝑝₃ ∙ sym 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) (sym 𝑝₂) + (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) + a₁₋₋ = solvePaths + + a₋₀₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₃ ∙ sym 𝑝₃) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₂ + a₋₀₋ = solvePaths + + a₋₁₋ : Square (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) 𝑝₁ + (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₁₋ = solvePaths + + a₋₋₀ : Square 𝑝₀ (sym 𝑝₂) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₁ + a₋₋₀ = solvePaths + + a₋₋₁ : Square (𝑝₂ ∙ 𝑝₃) (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) 𝑝₂ (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₋₁ = solvePaths + + -- this works but is slow (~2 min) + -- but resulting term is managable, and can be evaluated end typechecked quickly if imported in other module + + -- coh : Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + -- coh = solvePaths + +module CompCoherence (SA : NPath 7 A) where + open NPath SA + + LHS₀ RHS₀ : 𝑣₀ ≡ 𝑣₄ + LHS₀ = 𝑝₀ ∙∙ 𝑝₁ ∙ (𝑝₂ ∙ (𝑝₁ ∙ 𝑝₂) ⁻¹) ∙ 𝑝₁ ∙∙ 𝑝₂ ∙ 𝑝₃ + RHS₀ = 𝑝₀ ∙ (λ i → 𝑝₁ (i ∧ ~ i)) ∙ 𝑝₁ ∙ 𝑝₂ ∙ (λ i → 𝑝₂ (i ∨ ~ i)) ∙ 𝑝₃ + + LHS₁ RHS₁ : 𝑣₄ ≡ 𝑣₇ + LHS₁ = 𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆ + RHS₁ = 𝑝₄ ∙ refl ∙ 𝑝₅ ∙ refl ∙ refl ∙ 𝑝₆ + + LHS₀≡RHS₀ : LHS₀ ≡ RHS₀ + LHS₀≡RHS₀ = solvePaths + + LHS₁≡RHS₁ : LHS₁ ≡ RHS₁ + LHS₁≡RHS₁ = solvePaths + + LHS₀∙LHS₁≡RHS₀∙RHS₁ : LHS₀ ∙ LHS₁ ≡ RHS₀ ∙ RHS₁ + LHS₀∙LHS₁≡RHS₀∙RHS₁ = solvePaths + + -- _ : cong₂ _∙_ LHS₀≡RHS₀ LHS₁≡RHS₁ ≡ LHS₀∙LHS₁≡RHS₀∙RHS₁ + -- _ = solvePaths + + LHS₀⁻¹≡RHS₀⁻¹ : LHS₀ ⁻¹ ≡ RHS₀ ⁻¹ + LHS₀⁻¹≡RHS₀⁻¹ = solvePaths + + -- _ : cong (_⁻¹) LHS₀≡RHS₀ ≡ LHS₀⁻¹≡RHS₀⁻¹ + -- _ = solvePaths + + + +module GroupoidLaws (SA : NPath 6 A) where + open NPath SA + + symDist₅ : sym (𝑝₀ ∙ 𝑝₁ ∙ 𝑝₂ ∙ 𝑝₃ ∙ 𝑝₄ ∙ 𝑝₅) + ≡ sym 𝑝₅ ∙ sym 𝑝₄ ∙ sym 𝑝₃ ∙ sym 𝑝₂ ∙ sym 𝑝₁ ∙ sym 𝑝₀ + symDist₅ = solvePaths + + + module E₁ where + pa₀ pa₁ pa₂ pa₃ : 𝑣₀ ≡ 𝑣₆ + pa₀ = 𝑝₀ ∙ 𝑝₁ ∙ 𝑝₂ ∙ 𝑝₃ ∙ 𝑝₄ ∙ 𝑝₅ + pa₁ = ((((𝑝₀ ∙ 𝑝₁) ∙ 𝑝₂) ∙ 𝑝₃) ∙ 𝑝₄) ∙ 𝑝₅ + pa₂ = 𝑝₀ ∙ 𝑝₁ ∙' (refl ∙∙ 𝑝₂ ∙∙ (𝑝₃ ∙∙ 𝑝₄ ∙∙ 𝑝₅)) + pa₃ = 𝑝₀ ∙∙ 𝑝₁ ∙∙ (refl ∙' 𝑝₂ ∙ (𝑝₃ ∙' 𝑝₄ ∙ 𝑝₅)) + + assoc₅ : pa₀ ≡ pa₁ + assoc₅ = solvePaths + + pa₂≡pa₃ : pa₂ ≡ pa₃ + pa₂≡pa₃ = solvePaths + + pa₃≡pa₂ : pa₃ ≡ pa₂ + pa₃≡pa₂ = solvePaths + + sym[pa₃≡pa₂]≡pa₂≡pa₃ : sym (pa₃≡pa₂) ≡ pa₂≡pa₃ + sym[pa₃≡pa₂]≡pa₂≡pa₃ = refl + + pa₀≡pa₂ : pa₀ ≡ pa₂ + pa₀≡pa₂ = solvePaths + + pa₁≡pa₃ : pa₁ ≡ pa₃ + pa₁≡pa₃ = solvePaths + + pa₀≡pa₃ : pa₀ ≡ pa₃ + pa₀≡pa₃ = solvePaths + + + coherence : Square + assoc₅ pa₂≡pa₃ + pa₀≡pa₂ pa₁≡pa₃ + coherence = coh₃helper + + coh∙ : assoc₅ ∙ pa₁≡pa₃ ≡ pa₀≡pa₃ + coh∙ = comp-coh-helper + + + + + + + +module 2GroupoidLaws where + + module Triangle (SA : NPath 2 A) (X : A) where + open NPath SA + + + triangleIdentity : Cube + refl (assoc 𝑝₀ refl 𝑝₁) + (cong (𝑝₀ ∙_) (lUnit 𝑝₁)) (cong (_∙ 𝑝₁) (rUnit 𝑝₀)) + refl refl + triangleIdentity = solvePaths + + module Pentagon (SA : NPath 4 A) where + open NPath SA + + + pentagonIdentity' : assoc 𝑝₀ 𝑝₁ (𝑝₂ ∙ 𝑝₃) ∙ assoc (𝑝₀ ∙ 𝑝₁) 𝑝₂ 𝑝₃ + ≡ + cong (𝑝₀ ∙_) (assoc 𝑝₁ 𝑝₂ 𝑝₃) ∙∙ assoc 𝑝₀ (𝑝₁ ∙ 𝑝₂) 𝑝₃ ∙∙ cong (_∙ 𝑝₃) (assoc 𝑝₀ 𝑝₁ 𝑝₂) + pentagonIdentity' = solvePaths + + + -- solving this 4-cube commented code bellow takes + -- ~8 min (but memory spikes to more than 100GiB !) + + -- pentagonIdentity'≡pentagonIdentity : pentagonIdentity' ≡ pentagonIdentity 𝑝₀ 𝑝₁ 𝑝₂ 𝑝₃ + -- pentagonIdentity'≡pentagonIdentity = solvePaths + + module _ (f : A → B) where + + cf : ∀ {x y} → (p : x ≡ y) → f x ≡ f y + cf = cong f + + pentagonIdentityCong : + Square + (assoc (cf 𝑝₀) (cf 𝑝₁) ((cf 𝑝₂) ∙ (cf 𝑝₃)) ∙ assoc ((cf 𝑝₀) ∙ (cf 𝑝₁)) (cf 𝑝₂) (cf 𝑝₃)) + (sym (cong-∙ f _ _) ∙∙ cong cf (assoc 𝑝₀ (𝑝₁ ∙ 𝑝₂) 𝑝₃) ∙∙ cong-∙ f _ _) + ((cong (cf 𝑝₀ ∙_) (cong (cf 𝑝₁ ∙_) (sym (cong-∙ f _ _)) + ∙∙ sym (cong-∙ f _ _) + ∙∙ cong cf (assoc 𝑝₁ 𝑝₂ 𝑝₃)))) + ((cong (_∙ cf 𝑝₃) (cong (_∙ cf 𝑝₂) (sym (cong-∙ f _ _)) + ∙∙ sym (cong-∙ f _ _) + ∙∙ cong cf (sym (assoc 𝑝₀ 𝑝₁ 𝑝₂))))) + pentagonIdentityCong = solvePaths + + +module Glue (A B C D E : Type ℓ) + (e₀ : A ≃ B) (e₁ : B ≃ C) (e₂ : C ≃ D) (p : D ≡ E) where + + e0 : Square (ua e₀ ∙ ua e₁) (ua e₀ ∙∙ (ua e₁ ∙ ua e₂) ∙∙ p) refl (ua e₂ ∙ p) + e0 = solvePaths + + e0L : Square (cong List (cong List (ua e₀) ∙ cong List (ua e₁))) + (cong (List ∘S List) (ua e₀ ∙∙ ua e₁ ∙∙ ua e₂)) + refl (cong (List ∘S List) (ua e₂)) + e0L = solvePaths + + +module funTypes {x y z : A} (f : A → A → B) + (p : x ≡ y) + (q : y ≡ z) where + + + P Q : _≡_ {A = (A → B)} (λ x' → f x' x) (λ x' → f x' y) + P = (λ i x' → f x' (p i)) ∙∙ (λ i x' → f x' (q i)) ∙∙ (λ i x' → f x' (q (~ i))) + Q = refl ∙ (λ i x' → f x' (p i)) + + + + +module compPathR-PathP∙∙ + {x y : A} {p : x ≡ y} + where + + invSides-filler-rot' : (invSides-filler p p) ≡ (symP (invSides-filler (sym p) (sym p))) + + invSides-filler-rot' = solvePaths + + _ : invSides-filler-rot p ≡ invSides-filler-rot' + _ = solvePaths + + + + +module _ {A : Type} {x y z w v : A} (p' p'' : x ≡ y) (xr xr' : x ≡ x) (q : y ≡ z) (~r : w ≡ z) (r' r : z ≡ w) (s : w ≡ v) + (sq : Square xr (sym p'') p'' xr') where + + _ : refl ≡ λ i → p' (i ∨ ~ i) + _ = solvePaths + + _ : Cube + (λ i j → p' (i ∨ ~ i ∨ j ∨ ~ j)) (λ _ _ → y) + (λ _ _ → y) (λ _ _ → y) + (λ _ _ → y) (λ _ _ → y) + _ = solvePaths + + _ : (λ i → sq i (~ i)) ∙ refl ∙ ((λ i → sq (~ i) i) ∙ (λ i → sq i (~ i)) ∙' q ∙ sym (~r) ∙ + (~r ∙ (λ i → r (i ∧ ~ i)) ∙ + (r ∙ ((λ i → r (i ∨ ~ i))) ∙ s ))) + ≡ (λ i → sq i (~ i)) ∙ (q ∙ refl ∙ refl ∙ r ∙ s ∙ sym s) ∙ s + + + _ = solvePaths + + + +module _ {ℓ} where + + data D : Type ℓ where + x y z w : D + p : x ≡ y + q : y ≡ z + f : D → D + r : f z ≡ f w + + + _ : Square + (cong f (p ∙ q)) (cong f q ∙ r) + (cong f p) r + _ = solvePaths + + + +module CongCoherent (f : A → B) (SA : NPath 4 A) where + open NPath SA + + LHS RHS : 𝑣₀ ≡ 𝑣₄ + LHS = (𝑝₀ ∙ refl ∙ 𝑝₁) ∙ (𝑝₂ ∙ refl ∙ 𝑝₃) + RHS = 𝑝₀ ∙∙ (𝑝₁ ∙∙ refl ∙∙ 𝑝₂) ∙∙ 𝑝₃ + + solve[cong] cong[solve] : cong f LHS ≡ cong f RHS + + solve[cong] = solvePaths + + cong[solve] = cong (cong f) solvePaths + + _ : cong[solve] ≡ solve[cong] + _ = solvePaths diff --git a/Cubical/Tactics/PathSolver/NSolver/NSolver.agda b/Cubical/Tactics/PathSolver/NSolver/NSolver.agda new file mode 100644 index 0000000000..384145dfe0 --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/NSolver.agda @@ -0,0 +1,510 @@ +{- + + +This module implements a solver for computing fillings of n-dimensional cubes, assuming the boundary of the cube consists solely of paths. Each cell within the cubical complex must be expressible as a path applied to some interval expression. This path may be a diagonal or a face of a higher-dimensional cube, containing complex interval expressions. + +### Overview + +- **Assumptions**: + - Each cell in the complex can be expressed as a path with an interval expression. + - Paths can be complex, involving higher-dimensional cubes and intricate interval expressions. + +### Implementation + +- **Boundary Processing**: + - The initial step applies the generalised `cong-∙` lemma from `CongComp` to the entire boundary if necessary. + - Solver works by traversing the 1-dimensional skeleton (`1-skel`) of the boundary. + +- **Path Normalization**: + - For each vertex, the normal form of the path from the `i0ⁿ` corner to this vertex is computed. + - This normalization treats paths as elements of a free group, + to have robust(not necessary efficient) test for equality we use `unify` from `Agda.Builtin.Reflection`. + + +- **Filler Term Generation**: + - For vertices (becoming edges), the path lists are folded using the usual path composition operations. + - For edges (now becoming squares), the `compPath-filler` (specialized as `cpf`) is used when necessary + (depending of the comparison of length of normal form of path on the vertexes connected by edge). + +- **Generic Construction**: + - The algorithm supports arbitrary dimensions and can generate all coherence conditions for paths. + +### Core Definitions and Functions + +- **`solvePathsSolver`**: Main entry point for solving paths, managing reduction and boundary decomposition. +- **`markVert`**: Marks vertices with normal forms their associated paths, used during path traversal. +- **`foldBdTerm`**: Folds terms over the boundary, constructing the final term of solution. +- **notable helper Functions**: + - `isRedex?`, `η·`, `_ [·] _` for path composition and unification checks. + - `print[𝟚×]`, `printCellVerts` for debugging and visualization. + - `[𝟚×Term]→PathTerm`, `[𝟚×Term]→FillTerm` for generating cells in filler based on results of + computing normal forms during traversal + +### Examples and Usage + +The accompanying `Examples.agda` file demonstrates application of the solver and its limitations. + +-} + +{-# OPTIONS --safe -v testMarkVert:3 -v tactic:3 #-} + +module Cubical.Tactics.PathSolver.NSolver.NSolver where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ + +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) +open import Agda.Builtin.String +open import Agda.Builtin.Char + +open import Cubical.Data.Sigma.Base + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.PathSolver.Reflection +open import Cubical.Tactics.Reflection + +open import Cubical.Tactics.Reflection.Utilities + +-- open import Cubical.Tactics.PathSolver.Base +open import Cubical.Tactics.PathSolver.CongComp + +open import Cubical.Tactics.Reflection.QuoteCubical renaming (normaliseWithType to normaliseWithType') + +open import Cubical.Tactics.Reflection.Error +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.CuTerm +open import Cubical.Tactics.PathSolver.Reflection +open import Cubical.Tactics.Reflection.Variables +open import Cubical.Tactics.PathSolver.Degen + + +private + variable + ℓ : Level + A B : Type ℓ + +normaliseWithType : String → R.Type → R.Term → R.TC R.Term +normaliseWithType s ty tm = do + -- R.debugPrint "" 3 $ s <> " nwt: " ∷ₑ [ ty ]ₑ + normaliseWithType' ty tm + + +[𝟚×Term] : Type +[𝟚×Term] = List (Bool × R.Term) + +Vert : Type +Vert = List Bool + + +𝒏[_] : A → R.TC A +𝒏[_] = pure + + +isRedex? : (Bool × R.Term) → (Bool × R.Term) → R.TC Bool +isRedex? (b , x) (b' , x') = + if (b ⊕ b') + then + (((addNDimsToCtx 1 $ R.unify x x')>> pure true) + <|> pure false) + else (pure false) + +η· : Bool × R.Term → [𝟚×Term] → R.TC [𝟚×Term] +η· x [] = ⦇ [ ⦇ x ⦈ ] ⦈ +η· x (y ∷ xs) = do + b ← isRedex? x y + pure $ if b then xs else (x ∷ y ∷ xs) + +_[·]_ : [𝟚×Term] → [𝟚×Term] → R.TC [𝟚×Term] +xs [·] ys = foldrM η· ys xs + +invLi : [𝟚×Term] → [𝟚×Term] +invLi = L.map (λ (x , y) → not x , y) ∘S rev + + + +asPath : R.Term → R.TC (Maybe (Bool × R.Term)) +asPath tm = addNDimsToCtx 1 do + -- fi ← findInterval 1 <$> R.normalise tm + fi ← Mb.rec (pure nothing) (λ x → just <$> R.normalise x) $ findInterval 1 tm + + Mb.rec (⦇ nothing ⦈) (zz') fi + + where + zz : R.Term → R.TC (R.Term ⊎.⊎ Maybe (Bool × R.Term)) + zz (R.var zero []) = pure $ pure $ just (true , tm) + zz (R.def (quote ~_) v[ R.var zero [] ]) = pure $ pure (just (false , invVar zero tm)) + zz (R.con _ _) = pure $ pure nothing + zz (R.def (quote ~_) v[ R.var (suc k) [] ]) = + R.typeError ([ "imposible in asPath : ~ : " ]ₑ ++ₑ [ k ]ₑ) + zz tmI = pure (inl tmI) + + zz' : R.Term → R.TC (Maybe (Bool × R.Term)) + zz' = zz >=> + ⊎.rec (((extractIExprM >=& normIExpr) >=& IExpr→Term) >=> + (zz >=> ⊎.rec (λ tmI → + R.typeError ([ "imposible in asPath: " ]ₑ ++ₑ [ tm ]ₑ ++ₑ [ "\n\n" ]ₑ ++ₑ [ tmI ]ₑ)) + (pure))) pure + + +data CellVerts : Type where + cv0 : [𝟚×Term] → [𝟚×Term] → CellVerts + cvN : CellVerts → CellVerts → CellVerts + + +mapCellVerts : (f : [𝟚×Term] → [𝟚×Term]) → CellVerts → CellVerts +mapCellVerts f (cv0 x x₁) = cv0 (f x) (f x₁) +mapCellVerts f (cvN x x₁) = cvN (mapCellVerts f x) (mapCellVerts f x₁) + +mapCellVertsM : (f : [𝟚×Term] → R.TC [𝟚×Term]) → CellVerts → R.TC CellVerts +mapCellVertsM f (cv0 x x₁) = ⦇ cv0 (f x) (f x₁) ⦈ +mapCellVertsM f (cvN x x₁) = ⦇ cvN (mapCellVertsM f x) (mapCellVertsM f x₁) ⦈ + + +cellVert : CellVerts → Vert → R.TC [𝟚×Term] +cellVert (cv0 x x₂) (false ∷ []) = pure x +cellVert (cv0 x x₂) (true ∷ []) = pure x₂ +cellVert (cvN x x₂) (false ∷ x₃) = cellVert x x₃ +cellVert (cvN x x₂) (true ∷ x₃) = cellVert x₂ x₃ +cellVert _ _ = R.typeError $ [ "cellVert failed " ]ₑ + + + + +getAtomPa : R.Term → R.TC [𝟚×Term] +getAtomPa = (maybeToList <$>_) ∘S asPath + +print[𝟚×] : [𝟚×Term] → List R.ErrorPart +print[𝟚×] = + join ∘S (L.map (λ (b , t) + → ", (" ∷ₑ vlam "𝕚" t ∷ₑ [ ")" <> (if b then "" else "⁻¹") ]ₑ )) + +CellVerts→List : CellVerts → List (Vert × [𝟚×Term]) +CellVerts→List (cv0 x x₁) = ([ false ] , x) ∷ [ [ true ] , x₁ ] +CellVerts→List (cvN x x₁) = + L.map (λ (x , y) → (false ∷ x) , y) (CellVerts→List x) + ++ L.map ((λ (x , y) → true ∷ x , y)) (CellVerts→List x₁) + + +allEqual? : Discrete A → List A → Bool +allEqual? _≟_ (x ∷ (y ∷ xs)) = Dec→Bool (x ≟ y) and allEqual? _≟_ (y ∷ xs) +allEqual? _≟_ _ = true + + + +printCellVerts : CellVerts → List (R.ErrorPart) +printCellVerts = (join ∘ L.map + (λ (v , x) → L.map (if_then "□" else "◼") v ++ₑ print[𝟚×] x ++ₑ [ "\n" ]ₑ)) ∘ CellVerts→List + + + +module _ (ty : R.Type) where + getTermVerts : ℕ → R.Term → R.TC CellVerts + getTermVerts zero x₁ = R.typeError [ "imposible - getTermsVerts" ]ₑ + getTermVerts (suc zero) t = do + p ← getAtomPa t + pure (cv0 [] p) + + getTermVerts (suc n) t = do + gtv0 ← getTermVerts n (subfaceCell (just false ∷ repeat n nothing) t) + gtv1 ← getTermVerts n (subfaceCell (just true ∷ repeat n nothing) t) + + p0i ← getAtomPa (subfaceCell (nothing ∷ repeat n (just false)) t) + + cvN gtv0 <$> (mapCellVertsM (_[·] p0i) gtv1) + +getVert : ℕ → Vert → CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts)) → R.TC [𝟚×Term] +getVert zero v (hco xs _) = R.typeError [ "ran out of magic in getVert" ]ₑ +getVert (suc m) v (hco xs _) = do + (sf , x) ← Mb.rec (R.typeError [ "imposible getVert" ]ₑ) pure + $ findBy ((L.map just v ⊂?_) ∘ [_] ∘ fst) xs + let v' : Vert + v' = (L.map (snd) $ (filter ((λ { nothing → true ; _ → false }) ∘S fst) + (zipWith _,_ sf v))) + getVert m (true ∷ v') x +getVert _ x (cell' (_ , (_ , x₁)) _) = cellVert x₁ x + + +foldBdTermWithCuInput' = + let T = (CuTerm' ⊥ Unit × Maybe R.Term) + in List (ℕ × (T × T)) + + +foldBdTermWithCuInput = + let T = (CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts)) × Maybe R.Term) + in List (ℕ × (T × T)) + + + +module _ (ty : R.Type) where + + markVert : ℕ → ℕ → [𝟚×Term] → (CuTerm' ⊥ Unit) → R.TC (CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts))) + + getPaThruPartial : ℕ → Vert → List (SubFace × CuTerm' ⊥ Unit) → R.TC [𝟚×Term] + getPaThruPartial m v xs = do + (sf , x) ← Mb.rec (R.typeError [ "imposible gptp" ]ₑ) pure + $ findBy ((L.map just v ⊂?_) ∘ [_] ∘ fst) xs + let v' : Vert + v' = (L.map (snd) $ (filter ((λ { nothing → true ; _ → false }) ∘S fst) + (zipWith _,_ sf v))) + xs' ← markVert m (suc (sfDim sf)) [] x + p0 ← getVert m (false ∷ v') xs' + p1 ← getVert m (true ∷ v') xs' + p1 [·] (invLi p0) + + markVert zero dim w (hco x cu) = R.typeError [ "ran out of magic in markVert" ]ₑ + markVert (suc m) dim w (hco x cu) = do + -- markedVerts ← mapM (λ (sf , x) → ⦇ ⦇ sf ⦈ , markVert m (suc (sfDim sf)) [] x ⦈) x + paToLid ← invLi <$> (getPaThruPartial m (repeat dim false) x) + paToLid[·]w ← paToLid [·] w + markedCu ← markVert m dim (paToLid[·]w) cu + fixedVerts ← mapM (λ (sf , x) → do + vv ← (getVert m (L.map (Mb.fromJust-def false) sf) markedCu) + ⦇ ⦇ sf ⦈ , markVert m (suc (sfDim sf)) vv x ⦈) x + pure $ hco fixedVerts markedCu + markVert m dim w (cell x') = do + (mbX , x) ← UndegenCell.mbUndegen dim x' + -- R.debugPrint "testMarkVert" 0 $ "mv" ∷ₑ [ m ]ₑ + zz ← getTermVerts ty dim x >>= 𝒏[_] + -- ia ← getIArg dim x <|> + -- R.typeError (printCellVerts zz) + ia ← Mb.rec (⦇ nothing ⦈) ((extractIExprM >=> 𝒏[_]) >=& just) (findInterval dim x) + + zzT ← R.quoteTC zz + iaT ← R.quoteTC ia + + R.debugPrint "testMarkVert" 3 $ ("markVert : \n" ∷ₑ zzT ∷ₑ "\n" ∷ₑ [ iaT ]ₑ) + ⦇ cell' + ⦇ ⦇ mbX ⦈ , ⦇ ⦇ ia ⦈ , mapCellVertsM (_[·] w) zz ⦈ ⦈ + ⦇ x ⦈ + ⦈ + + markVertSnd : ℕ → ℕ → [𝟚×Term] → ((CuTerm' ⊥ Unit) × A) + → R.TC (CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts)) × A) + markVertSnd n m tms (x , y) = ⦇ markVert n m tms x , ⦇ y ⦈ ⦈ + + markVertBd : foldBdTermWithCuInput' + → R.TC foldBdTermWithCuInput + markVertBd [] = R.typeError [ "markVertBd undefined" ]ₑ + markVertBd (_ ∷ []) = R.typeError [ "markVertBd undefined" ]ₑ + markVertBd xs = do + let dim = predℕ (length xs) + v0 = repeat dim false + fcs0 ← mapM (λ (k , (c0 , _ )) → + do R.debugPrint "solvePaths" 0 $ "solvePaths - markVert dim: " ∷ₑ [ k ]ₑ + markVertSnd 100 dim [] c0 <|> fail ("fcs0 fail" ∷nl [ k ]ₑ)) xs + fcs0₀ ← Mb.rec (R.typeError [ "imposible" ]ₑ) + (λ y → mapM (λ k → (getVert 100 (replaceAt k true v0)) (fst y) + <|> fail ("fcs1 fail" ∷nl [ k ]ₑ)) (range dim)) + (lookupMb fcs0 0) + fcs0₁ ← Mb.rec (R.typeError [ "imposible" ]ₑ) + (getVert 100 (replaceAt (predℕ dim) true v0) ∘S fst) (lookupMb fcs0 1) + + fcs1 ← mapM (idfun _) + (zipWith (markVertSnd 100 dim) (fcs0₁ ∷ fcs0₀) ((snd ∘S snd) <$> xs)) + pure $ zipWithIndex (zipWith _,_ fcs0 fcs1) + + + +flipOnFalse : Bool → R.Term → R.Term +flipOnFalse b t = if b then t else R.def (quote ~_) v[ t ] + + + +cpf : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ j → x ≡ q j) p (p ∙ q) + +cpf {x = x} {y} p q i z = hcomp + (λ { j (z = i1) → q (i ∧ j) + ; j (z = i0) → x + ; j (i = i0) → p z + }) + (p z) + +[𝟚×Term]→PathTerm : [𝟚×Term] → R.Term +[𝟚×Term]→PathTerm [] = q[ refl ] +[𝟚×Term]→PathTerm ((b , tm) ∷ []) = + R∙ (vlam "_" (liftVars (subfaceCell [ just (not b) ] tm))) + (vlam "𝕚'" (if b then tm else (invVar zero tm))) --(if b then tm else Rsym tm) +[𝟚×Term]→PathTerm ((b , tm) ∷ xs) = R∙ ([𝟚×Term]→PathTerm xs) + (vlam "𝕚'" (if b then tm else (invVar zero tm))) + +[𝟚×Term]→FillTerm : Bool × R.Term → [𝟚×Term] → R.Term +[𝟚×Term]→FillTerm (b , tm) [] = + R.def (quote cpf) ((vlam "_" (liftVars (subfaceCell [ just (not b) ] tm))) + v∷ v[ (vlam "𝕚'" (if b then tm else (invVar zero tm))) ]) +[𝟚×Term]→FillTerm (b , tm) xs = + R.def (quote cpf) ([𝟚×Term]→PathTerm xs v∷ + v[ (vlam "𝕚'" (if b then tm else (invVar zero tm))) ]) + +dbgId : ∀ {ℓ} {A : Type ℓ} → String → A → A +dbgId _ x = x + +module MakeFoldTerm (t0 : R.Term) where + + + cellTerm : ℕ → (Maybe IExpr) × ((Maybe (Bool × R.Term) × [𝟚×Term])) → R.Term → R.Term + -- cellTerm = {!!} + cellTerm dim (mbi , nothing , []) t = + (liftVars t) + cellTerm dim (mbi , nothing , tl@(_ ∷ _)) _ = --R.unknown + R.def (quote $≡) (liftVarsFrom (suc dim) 0 ([𝟚×Term]→PathTerm tl) v∷ + v[ R.def (quote ~_) v[ 𝒗 dim ] ]) + cellTerm dim (just ie , just (b , tm) , tl) _ = --vlamⁿ 1 (liftVarsFrom 1 0 t) + + R.def (quote $≡) + ((R.def (quote $≡) (liftVarsFrom (suc dim) 0 ([𝟚×Term]→FillTerm (b , tm) tl) v∷ + -- v[ (IExpr→Term ie) ]) v∷ + v[ flipOnFalse (b) (IExpr→Term ie) ]) v∷ + v[ R.def (quote ~_) v[ 𝒗 dim ] ])) + cellTerm _ _ _ = R.lit (R.string ("unexpected in MakeFoldTerm.cellTerm")) + + + ctils : List (SubFace × (CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts)))) → + R.TC (List (SubFace × CuTerm)) + + ctil : ℕ → (CuTerm' ⊥ (Maybe (R.Term × R.Term) × ((Maybe IExpr) × CellVerts))) → R.TC CuTerm + ctil dim (hco x c) = + ⦇ hco ⦇ pure (repeat dim nothing ++ [ just true ] , cell + -- (R.def (quote dbgId) (R.lit (R.string "ctill-fill") v∷ v[ t0 ]) ) + (liftVarsFrom (suc dim) zero t0) + ) + ∷ + ctils x ⦈ + (ctil dim c) ⦈ + ctil dim (cell' (mbt , cv) x) = cell' tt <$> + let ct = (cellTerm dim (fst cv , cellVertsHead (snd cv)) x) + in Mb.rec + -- (pure $ R.def (quote dbgId) (R.lit (R.string "ctil") v∷ v[ ct ]) ) + (pure ct) + (λ tmUDG → UndegenCell.undegenCell dim tmUDG ct + ) mbt + + where + cellVertsHead : CellVerts → Maybe (Bool × R.Term) × [𝟚×Term] + cellVertsHead cv = + let l = L.map (snd) $ CellVerts→List cv + lM = L.map (length) l + + + in if (allEqual? discreteℕ lM) then nothing , Mb.fromJust-def [] (listToMaybe l) else + let maxL = foldr max 0 lM + in Mb.rec (nothing , []) (λ x → listToMaybe x , tail x) (findBy (λ x → Dec→Bool $ discreteℕ (length x) maxL ) l) + + + ctils [] = ⦇ [] ⦈ + ctils ((sf , x) ∷ xs) = + ⦇ ⦇ pure (sf ++ [ nothing ]) , ctil (suc (sfDim sf)) x ⦈ ∷ ctils xs ⦈ + + + + +foldBdTerm : R.Type → R.Term → foldBdTermWithCuInput + → R.TC R.Term +foldBdTerm _ _ [] = R.typeError [ "foldBdTerm undefined for 0 dim" ]ₑ +foldBdTerm ty f0 xs = do + let dim = length xs + needsCongFill = any? (L.map (λ { (_ , ((_ , nothing) , (_ , nothing))) → false ; _ → true} ) xs) + t0UL ← normaliseWithType "mkFoldTerm" ty + (subfaceCell (repeat (predℕ dim) (just false)) f0) + let t0 = liftVarsFrom dim zero t0UL + toTerm {A = Unit} dim <$> + (⦇ hco + (mapM (idfun _) $ join $ zipWith + (λ k (_ , (cu0 , cu1)) → + let sf0 = replaceAt k (just false) (repeat dim nothing) + sf1 = replaceAt k (just true) (repeat dim nothing) + prmV = invVar 0 ∘S remapVars (λ k → + if (k <ℕ dim) then (if (k =ℕ (predℕ dim)) then zero else suc k) + else k) + + fc : SubFace → + (CuTerm' ⊥ (Maybe (R.Term × R.Term) × Maybe IExpr × CellVerts) × + Maybe R.Term) → + List _ + fc sf cu = + let cuTm' = ((prmV ∘S ToTerm.toTerm (defaultCtx dim)) <$> + MakeFoldTerm.ctil t0UL (predℕ dim) (fst cu)) + cuTm = ⦇ cell cuTm' ⦈ + in [ ((sf ,_)) <$> + (if (not needsCongFill) + then cuTm + else do + cpa ← cell <$> + (Mb.rec (subfaceCellNoDrop (just true ∷ repeat (predℕ dim) nothing) <$> cuTm') + (λ pa → pure $ (prmV pa)) (snd cu)) + ⦇ hco + (pure ( (just true ∷ repeat (predℕ dim) nothing , cpa) + ∷ [ just false ∷ repeat (predℕ dim) nothing , + cell t0 ])) + cuTm ⦈) ] + + in fc sf0 cu0 ++ fc sf1 cu1) + (range dim) xs ) + ⦇ cell ⦇ t0 ⦈ ⦈ ⦈ ) -- >>= normaliseCells dim) + + + +doNotReduceInPathSolver = [ quote ua ] + + + +toNoCons : ℕ → CuTerm → R.TC (CuTerm' ⊥ Unit × Maybe R.Term) +toNoCons dim cu = + Mb.rec + (do ptm ← addNDimsToCtx (suc dim) $ R.normalise $ (ToTerm.toTerm (defaultCtx (suc dim)) (fillCongs 100 dim cu)) + pure $ appCongs dim cu , just ptm) + (λ x → ⦇ ⦇ x ⦈ , ⦇ nothing ⦈ ⦈) + (tryCastAsNoCong cu) + + +solvePathsSolver : R.Type → R.TC R.Term +solvePathsSolver goal = R.withReduceDefs (false , doNotReduceInPathSolver) do + R.debugPrint "solvePaths" 0 $ [ "solvePaths - start" ]ₑ + hTy ← wait-for-term goal >>= + (λ x → (R.debugPrint "solvePaths" 0 $ "solvePaths - " ∷ₑ [ x ]ₑ) >> pure x) + >>= (R.normalise <|>>= λ x → R.typeError $ "failed to normalise goal" ∷nl [ x ]ₑ ) + bdTM@(A , fcs) ← matchNCube hTy + R.debugPrint "solvePaths" 0 $ [ "solvePaths - matchNCube done" ]ₑ + let dim = length fcs + + (t0 , _) ← Mb.rec (R.typeError [ "imposible in solvePaths''" ]ₑ) pure (lookupMb fcs zero) + + cuFcs ← ((zipWithIndex <$> quoteBd bdTM + -- <|> R.typeError [ "quoteBd - failed" ]ₑ + ) >>= mapM + (λ (k , (cu0 , cu1)) → + (R.debugPrint "solvePaths" 0 $ "solvePaths - solve cong dimension : " ∷ₑ [ k ]ₑ) + >> ⦇ ⦇ k ⦈ , ⦇ toNoCons (predℕ dim) cu0 , toNoCons (predℕ dim) cu1 ⦈ ⦈ <|> + R.typeError [ "toNoCons - failed" ]ₑ)) + + + markVertBd A cuFcs >>= foldBdTerm A t0 + + +macro + + solvePaths : R.Term → R.TC Unit + solvePaths h = do + solution ← + (R.inferType h >>= solvePathsSolver) + R.unify solution h <|> R.typeError ("unify - failed:" ∷nl [ solution ]ₑ ) + + infixr 2 solvePathsTest_ + + solvePathsTest_ : R.Term → R.Term → R.TC Unit + solvePathsTest_ goal h = assertNoErr h ( + do solution ← solvePathsSolver goal + R.checkType solution goal) diff --git a/Cubical/Tactics/PathSolver/NSolver/Tests.agda b/Cubical/Tactics/PathSolver/NSolver/Tests.agda new file mode 100644 index 0000000000..e824560e82 --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Tests.agda @@ -0,0 +1,43 @@ +{- + +This module contains a set of test cases designed to ensure the robustness of the `NSolver` module. The multitude of test cases is motivated by the necessity to cover various aspects that impact the solver's correctness. + +### Aspects Covered + +- **Dimensionality**: + - Goals of more than 2 dimensions are treated slightly diferent in some situations, also some edgecases can + apear only in higher dimensions. +- **Functoriality of Congruence**: + - Verifies that solving paths may require applying functoriality of congruence. +- **Path Definition Scope**: + - Tests paths defined as free variables, higher constructors, and abstract definitions. +- **Degenerated Interval Expressions**: + - Verifies handling of degenerated interval expressions present in paths. + +### Importance of Coverage + +Covering these aspects in different combinations is crucial for maintaining the solver's code. Experience during development has shown that some errors can be specific to arbitrary combinations of the above aspects. Having the ability to quickly receive feedback about the precise combination of factors triggering an error is indispensable when introducing optimizations and heuristics in the solver. + +### Test Modules + +- **Const.agda**: + - Tests solver on basic properties of constant paths. +- **GroupoidLaws.agda**: + - Checks the solver against groupoid laws. +- **Cong.agda**: + - Verifies that the solver correctly handles congruence properties. +- **Coherence.agda**: + - Tests sollver on coherence of its own solutions. + +-} + +{-# OPTIONS --safe -v 0 #-} + +module Cubical.Tactics.PathSolver.NSolver.Tests where + + +import Cubical.Tactics.PathSolver.NSolver.Tests.Const +import Cubical.Tactics.PathSolver.NSolver.Tests.GroupoidLaws +import Cubical.Tactics.PathSolver.NSolver.Tests.Cong +import Cubical.Tactics.PathSolver.NSolver.Tests.Coherence + diff --git a/Cubical/Tactics/PathSolver/NSolver/Tests/Coherence.agda b/Cubical/Tactics/PathSolver/NSolver/Tests/Coherence.agda new file mode 100644 index 0000000000..55c813381f --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Tests/Coherence.agda @@ -0,0 +1,77 @@ +{-# OPTIONS --safe -v 0 #-} + +module Cubical.Tactics.PathSolver.NSolver.Tests.Coherence where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Tactics.PathSolver.NSolver.NSolver +open import Cubical.Tactics.PathSolver.Path +open import Cubical.Tactics.Reflection.Error + + +private + variable + ℓ ℓ' : Level + + +module CompCoherence {A : Type ℓ} (SA : NPath 7 A) where + open NPath SA + + module Basic where + LHS₀ RHS₀ : 𝑣₀ ≡ 𝑣₂ + LHS₀ = 𝑝₀ ∙∙ 𝑝₁ ∙∙ refl + RHS₀ = 𝑝₀ ∙∙ refl ∙∙ 𝑝₁ + + LHS₁ RHS₁ : 𝑣₂ ≡ 𝑣₃ + LHS₁ = 𝑝₂ + RHS₁ = 𝑝₂ + + LHS₀≡RHS₀ : LHS₀ ≡ RHS₀ + LHS₀≡RHS₀ = solvePaths + + LHS₁≡RHS₁ : LHS₁ ≡ RHS₁ + LHS₁≡RHS₁ = solvePaths + + LHS₀∙LHS₁≡RHS₀∙RHS₁ : LHS₀ ∙ LHS₁ ≡ RHS₀ ∙ RHS₁ + LHS₀∙LHS₁≡RHS₀∙RHS₁ = solvePaths + + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong₂ _∙_ LHS₀≡RHS₀ LHS₁≡RHS₁ ≡ LHS₀∙LHS₁≡RHS₀∙RHS₁ + + LHS₀⁻¹≡RHS₀⁻¹ : LHS₀ ⁻¹ ≡ RHS₀ ⁻¹ + LHS₀⁻¹≡RHS₀⁻¹ = solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong (_⁻¹) LHS₀≡RHS₀ ≡ LHS₀⁻¹≡RHS₀⁻¹ + + module WithDegens where + LHS₀ RHS₀ : 𝑣₀ ≡ 𝑣₄ + LHS₀ = 𝑝₀ ∙∙ 𝑝₁ ∙ (𝑝₂ ∙ (𝑝₁ ∙ 𝑝₂) ⁻¹) ∙ 𝑝₁ ∙∙ 𝑝₂ ∙ 𝑝₃ + RHS₀ = 𝑝₀ ∙ (λ i → 𝑝₁ (i ∧ ~ i)) ∙ 𝑝₁ ∙ 𝑝₂ ∙ (λ i → 𝑝₂ (i ∨ ~ i)) ∙ 𝑝₃ + + LHS₁ RHS₁ : 𝑣₄ ≡ 𝑣₇ + LHS₁ = 𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆ + RHS₁ = 𝑝₄ ∙ refl ∙ 𝑝₅ ∙ refl ∙ refl ∙ 𝑝₆ + + LHS₀≡RHS₀ : LHS₀ ≡ RHS₀ + LHS₀≡RHS₀ = solvePaths + + LHS₁≡RHS₁ : LHS₁ ≡ RHS₁ + LHS₁≡RHS₁ = solvePaths + + LHS₀∙LHS₁≡RHS₀∙RHS₁ : LHS₀ ∙ LHS₁ ≡ RHS₀ ∙ RHS₁ + LHS₀∙LHS₁≡RHS₀∙RHS₁ = solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong₂ _∙_ LHS₀≡RHS₀ LHS₁≡RHS₁ ≡ LHS₀∙LHS₁≡RHS₀∙RHS₁ + LHS₀⁻¹≡RHS₀⁻¹ : LHS₀ ⁻¹ ≡ RHS₀ ⁻¹ + LHS₀⁻¹≡RHS₀⁻¹ = solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong (_⁻¹) LHS₀≡RHS₀ ≡ LHS₀⁻¹≡RHS₀⁻¹ diff --git a/Cubical/Tactics/PathSolver/NSolver/Tests/Cong.agda b/Cubical/Tactics/PathSolver/NSolver/Tests/Cong.agda new file mode 100644 index 0000000000..25f0461a04 --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Tests/Cong.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --safe -v 0 #-} + +module Cubical.Tactics.PathSolver.NSolver.Tests.Cong where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Tactics.PathSolver.NSolver.NSolver +open import Cubical.Tactics.PathSolver.Path +open import Cubical.Tactics.Reflection.Error + + +private + variable + ℓ ℓ' : Level + + + +module Refl {A : Type ℓ} {B : Type ℓ'} (f : A → B) (a : A) where + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong f (refl {x = a} ∙ refl) ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong f (refl ∙ (refl {x = a} ∙ refl) ∙ refl) ∙ cong f ((refl ∙ refl) ∙ refl) ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + Square + ((cong f (((refl {x = a} ∙ refl) ∙ refl) ∙ refl) ∙ refl) ∙ refl) + refl + (refl ∙ cong f (refl ∙ refl ∙ refl) ∙ cong f (refl ∙ refl)) + (cong f ((refl ∙ refl) ∙∙ (refl ∙ refl) ∙∙ (refl ∙ refl ))) + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (congP (λ _ → cong f) (assoc (refl {x = a}) refl refl)) + (cong (refl ∙_) (lUnit refl) ∙ solvePaths) + (cong (_∙ refl) (rUnit refl) ∙ solvePaths) + refl refl + +module CongCoherent {A : Type ℓ} {B : Type ℓ'} (f : A → B) (SA : NPath 4 A) where + open NPath SA + + LHS RHS : 𝑣₀ ≡ 𝑣₄ + LHS = (𝑝₀ ∙ refl ∙ 𝑝₁) ∙ (𝑝₂ ∙ refl ∙ 𝑝₃) + RHS = 𝑝₀ ∙∙ (𝑝₁ ∙∙ refl ∙∙ 𝑝₂) ∙∙ 𝑝₃ + + solve[cong] cong[solve] : cong f LHS ≡ cong f RHS + + solve[cong] = solvePaths + + cong[solve] = cong (cong f) solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + cong[solve] ≡ solve[cong] diff --git a/Cubical/Tactics/PathSolver/NSolver/Tests/Const.agda b/Cubical/Tactics/PathSolver/NSolver/Tests/Const.agda new file mode 100644 index 0000000000..f6718a3441 --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Tests/Const.agda @@ -0,0 +1,101 @@ +{-# OPTIONS --safe -v 0 #-} + +module Cubical.Tactics.PathSolver.NSolver.Tests.Const where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Tactics.PathSolver.NSolver.NSolver +open import Cubical.Tactics.Reflection.Error + +private + variable + ℓ ℓ' : Level + +module Var {A : Type ℓ} (a : A) where + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl {x = a} ∙ refl ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl ∙ (refl {x = a} ∙ refl) ∙ refl ∙ (refl ∙ refl) ∙ refl ≡ refl + _ : ResultIs ✓-pass + _ = solvePathsTest + Square + (((((refl {x = a} ∙ refl) ∙ refl) ∙ refl) ∙ refl) ∙ refl) + refl + (refl ∙ refl ∙ refl ∙ refl ∙ refl ∙ refl) + ((refl ∙ refl) ∙∙ (refl ∙ refl) ∙∙ (refl ∙ refl )) + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (assoc (refl {x = a}) refl refl) + (cong (refl ∙_) (lUnit refl)) (cong (_∙ refl) (rUnit refl)) + refl refl + + + module Def where + abstract + a' : A + a' = a + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl {x = a'} ∙ refl ≡ refl + + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl ∙ (refl {x = a'} ∙ refl) ∙ refl ∙ (refl ∙ refl) ∙ refl ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + Square + (((((refl {x = a'} ∙ refl) ∙ refl) ∙ refl) ∙ refl) ∙ refl) + refl + (refl ∙ refl ∙ refl ∙ refl ∙ refl ∙ refl) + ((refl ∙ refl) ∙∙ (refl ∙ refl) ∙∙ (refl ∙ refl )) + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (assoc (refl {x = a'}) refl refl) + (cong (refl ∙_) (lUnit refl)) (cong (_∙ refl) (rUnit refl)) + refl refl + + + +module DataType {ℓ} where + + data A : Type ℓ where + a : A + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl {x = a} ∙ refl ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + refl ∙ (refl {x = a} ∙ refl) ∙ refl ∙ (refl ∙ refl) ∙ refl ≡ refl + + _ : ResultIs ✓-pass + _ = solvePathsTest + Square + (((((refl {x = a} ∙ refl) ∙ refl) ∙ refl) ∙ refl) ∙ refl) + refl + (refl ∙ refl ∙ refl ∙ refl ∙ refl ∙ refl) + ((refl ∙ refl) ∙∙ (refl ∙ refl) ∙∙ (refl ∙ refl )) + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (assoc (refl {x = a}) refl refl) + (cong (refl ∙_) (lUnit refl)) (cong (_∙ refl) (rUnit refl)) + refl refl + diff --git a/Cubical/Tactics/PathSolver/NSolver/Tests/GroupoidLaws.agda b/Cubical/Tactics/PathSolver/NSolver/Tests/GroupoidLaws.agda new file mode 100644 index 0000000000..64846ebc6e --- /dev/null +++ b/Cubical/Tactics/PathSolver/NSolver/Tests/GroupoidLaws.agda @@ -0,0 +1,274 @@ +{-# OPTIONS --safe -v 0 #-} + +module Cubical.Tactics.PathSolver.NSolver.Tests.GroupoidLaws where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Tactics.PathSolver.NSolver.NSolver +open import Cubical.Tactics.Reflection.Error + + +private + variable + ℓ ℓ' : Level + + +module Ω-Tests where + module Var (A : Type ℓ) (a : A) (p : a ≡ a) where + _ : ResultIs ✓-pass + _ = solvePathsTest + p ∙ p ∙ p ∙ p ∙ p ≡ ((((p ∙ p) ∙ p) ∙ p) ∙ p) + + + _ : ResultIs ✓-pass + _ = solvePathsTest + p ∙ refl ∙ p ∙ refl ∙ p ∙ refl ∙ refl ∙ p ∙ refl ∙ refl ∙ p ∙ refl + ≡ p ∙ p ∙ p ∙ p ∙ p + + + _ : ResultIs ✓-pass + _ = solvePathsTest + p ∙ p ⁻¹ ∙ p ∙' p ∙ p ⁻¹ ∙ p ∙ p ∙ p ⁻¹ ∙ p ⁻¹ ∙ p ⁻¹ ≡ refl + + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (assoc p refl p) + (cong (p ∙_) (lUnit p)) (cong (_∙ p) (rUnit p)) + refl refl + + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + (λ i j → p (i ∨ ~ i ∨ j ∨ ~ j)) (λ _ _ → a) + (λ _ _ → a) (λ _ _ → a) + (λ _ _ → a) (λ _ _ → a) + + + + module HIT where + open import Cubical.HITs.S1.Base + + _ : ResultIs ✓-pass + _ = solvePathsTest + loop ∙ loop ∙ loop ∙ loop ∙ loop ≡ ((((loop ∙ loop) ∙ loop) ∙ loop) ∙ loop) + + + _ : ResultIs ✓-pass + _ = solvePathsTest + loop ∙ refl ∙ loop ∙ refl ∙ loop ∙ refl ∙ refl ∙ loop ∙ refl ∙ refl ∙ loop ∙ refl + ≡ loop ∙ loop ∙ loop ∙ loop ∙ loop + + + _ : ResultIs ✓-pass + _ = solvePathsTest + loop ∙ loop ⁻¹ ∙ loop ∙' loop ∙ loop ⁻¹ ∙ loop ∙ loop ∙ loop ⁻¹ ∙ loop ⁻¹ ∙ loop ⁻¹ ≡ refl + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube + refl (assoc loop refl loop) + (cong (loop ∙_) (lUnit loop)) (cong (_∙ loop) (rUnit loop)) + refl refl + + + + +module NoCong where + module Var (A : Type ℓ) (a₀ a₁ a₂ a₃ a₄ a₅ a₆ a₇ : A) + (𝑝₀ : a₀ ≡ a₁) + (𝑝₁ : a₁ ≡ a₂) + (𝑝₂ : a₂ ≡ a₃) + (𝑝₃ : a₃ ≡ a₄) + (𝑝₄ : a₄ ≡ a₅) + (𝑝₅ : a₅ ≡ a₆) + (𝑝₆ : a₆ ≡ a₇) where + + a₀₋₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) 𝑝₀ (𝑝₂ ∙ 𝑝₃) + a₀₋₋ = solvePaths + + a₁₋₋ : Square (𝑝₃ ∙ sym 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) (sym 𝑝₂) + (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) + a₁₋₋ = solvePaths + + a₋₀₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₃ ∙ sym 𝑝₃) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₂ + a₋₀₋ = solvePaths + + a₋₁₋ : Square (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) 𝑝₁ + (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₁₋ = solvePaths + + a₋₋₀ : Square 𝑝₀ (sym 𝑝₂) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₁ + a₋₋₀ = solvePaths + + a₋₋₁ : Square (𝑝₂ ∙ 𝑝₃) (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) 𝑝₂ (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₋₁ = solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + +module HIT {ℓ} where + + + data A : Type ℓ where + a₀ a₁ a₂ a₃ a₄ a₅ a₆ a₇ : A + 𝑝₀ : a₀ ≡ a₁ + 𝑝₁ : a₁ ≡ a₂ + 𝑝₂ : a₂ ≡ a₃ + 𝑝₃ : a₃ ≡ a₄ + 𝑝₄ : a₄ ≡ a₅ + 𝑝₅ : a₅ ≡ a₆ + 𝑝₆ : a₆ ≡ a₇ + + a₀₋₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) 𝑝₀ (𝑝₂ ∙ 𝑝₃) + a₀₋₋ = solvePaths + + a₁₋₋ : Square (𝑝₃ ∙ sym 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) (sym 𝑝₂) + (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) + a₁₋₋ = solvePaths + + a₋₀₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₃ ∙ sym 𝑝₃) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₂ + a₋₀₋ = solvePaths + + a₋₁₋ : Square (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) 𝑝₁ + (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₁₋ = solvePaths + + a₋₋₀ : Square 𝑝₀ (sym 𝑝₂) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₁ + a₋₋₀ = solvePaths + + a₋₋₁ : Square (𝑝₂ ∙ 𝑝₃) (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) 𝑝₂ (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₋₁ = solvePaths + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + +mkPa : ∀ {ℓ} {A : Type ℓ} → (𝒑 : I → A) → 𝒑 i0 ≡ 𝒑 i1 +mkPa 𝒑 i = 𝒑 i + +module Edges&Diags {ℓ} (A : Type ℓ) + (a⁵ : I → I → I → I → I → A) where + + + 𝑝₀ : _ ≡ _ + 𝑝₀ = mkPa λ i → a⁵ i0 i i0 i (~ i) + + 𝑝₁ : _ ≡ _ + 𝑝₁ = mkPa λ i → a⁵ i i1 i i1 i0 + + 𝑝₂ : _ ≡ _ + 𝑝₂ = mkPa λ i → a⁵ i1 (~ i) i1 i1 i0 + + 𝑝₃ : _ ≡ _ + 𝑝₃ = mkPa λ i → a⁵ (~ i) i (~ i) (~ i) i + + 𝑝₄ : _ ≡ _ + 𝑝₄ = mkPa λ _ → a⁵ i0 i1 i0 i0 i1 + + 𝑝₅ : _ ≡ _ + 𝑝₅ = mkPa λ i → a⁵ (i ∧ ~ i) i1 i0 i0 (i ∨ ~ i) + + 𝑝₆ : _ ≡ _ + 𝑝₆ = mkPa λ i → a⁵ i0 i1 i0 i0 (~ i) + + a₀₋₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) 𝑝₀ (𝑝₂ ∙ 𝑝₃) + a₀₋₋ = solvePaths + + a₁₋₋ : Square (𝑝₃ ∙ sym 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) (sym 𝑝₂) + (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) + a₁₋₋ = solvePaths + + a₋₀₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₃ ∙ sym 𝑝₃) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₂ + a₋₀₋ = solvePaths + + a₋₁₋ : Square (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) 𝑝₁ + (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₁₋ = solvePaths + + a₋₋₀ : Square 𝑝₀ (sym 𝑝₂) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₁ + a₋₋₀ = solvePaths + + a₋₋₁ : Square (𝑝₂ ∙ 𝑝₃) (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) 𝑝₂ (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₋₁ = solvePaths + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + + +module InSubTerms {ℓ} (A : Type ℓ) + (a₀ a₁ a₂ a₃ : A) + (p₀₁ : a₀ ≡ a₁) + (p₁₂ : a₁ ≡ a₂) + + (f : A → I → A) + (g : A → A → A → A) + (h : g a₀ a₁ ≡ g (f a₂ i0) a₃) + (l : g (f a₂ i1) a₃ (f a₀ i1) ≡ a₀) where + + + 𝑝₀ : _ ≡ _ + 𝑝₀ = mkPa λ i → g (p₀₁ i) a₀ (f a₁ i) + + 𝑝₁ : _ ≡ _ + 𝑝₁ = mkPa λ i → g (p₀₁ (~ i)) (p₀₁ i) (f (p₀₁ (~ i)) i1) + + 𝑝₂ : _ ≡ _ + 𝑝₂ = mkPa λ i → h i (f a₀ i1) + + 𝑝₃ : _ ≡ _ + 𝑝₃ = mkPa λ i → g (f a₂ i) a₃ (f a₀ i1) + + 𝑝₄ : _ ≡ _ + 𝑝₄ = l + + 𝑝₅ : _ ≡ _ + 𝑝₅ = p₀₁ + + 𝑝₆ : _ ≡ _ + 𝑝₆ = p₁₂ + + + a₀₋₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) 𝑝₀ (𝑝₂ ∙ 𝑝₃) + a₀₋₋ = solvePaths + + a₁₋₋ : Square (𝑝₃ ∙ sym 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) (sym 𝑝₂) + (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) + a₁₋₋ = solvePaths + + a₋₀₋ : Square (𝑝₀ ∙ 𝑝₁) (𝑝₃ ∙ sym 𝑝₃) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₂ + a₋₀₋ = solvePaths + + a₋₁₋ : Square (𝑝₁ ∙∙ 𝑝₂ ∙∙ 𝑝₃) (𝑝₂ ∙ 𝑝₃ ∙ (𝑝₄ ∙∙ 𝑝₅ ∙∙ 𝑝₆)) 𝑝₁ + (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₁₋ = solvePaths + + a₋₋₀ : Square 𝑝₀ (sym 𝑝₂) (𝑝₀ ∙∙ 𝑝₁ ∙∙ 𝑝₂) 𝑝₁ + a₋₋₀ = solvePaths + + a₋₋₁ : Square (𝑝₂ ∙ 𝑝₃) (((𝑝₃ ∙' 𝑝₄) ∙' 𝑝₅) ∙' 𝑝₆) 𝑝₂ (𝑝₄ ∙ 𝑝₅ ∙ 𝑝₆) + a₋₋₁ = solvePaths + + + _ : ResultIs ✓-pass + _ = solvePathsTest + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + + +module FromI where + module Var (A : Type ℓ) (aᵢ : I → A) + (a₀ a₁ : A) + (p : a₀ ≡ aᵢ i0)(q : aᵢ i1 ≡ a₁) where + + _ : ResultIs ⊘-fail + _ = solvePathsTest + (p ∙ λ i → aᵢ i) ∙ q ≡ p ∙ ((λ i → aᵢ i) ∙ q) diff --git a/Cubical/Tactics/PathSolver/Path.agda b/Cubical/Tactics/PathSolver/Path.agda new file mode 100644 index 0000000000..620d77b768 --- /dev/null +++ b/Cubical/Tactics/PathSolver/Path.agda @@ -0,0 +1,441 @@ +{- +This module defines the `NPath` type, which conveniently represents a sequence of paths in a given type `A`. +This abstraction is primarily intended to ease the introduction of multiple paths into the context, +facilitating the creation of tests and examples. + +The module also provides several utility lemmas for composing squares and cubes, which are frequently used in the accompanying solvers within the `PathSolver` module. +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.Path where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Nat.Order.Recursive + +open import Cubical.Data.Sigma.Base + + + + +record NPath {ℓ} n (A : Type ℓ) : Type ℓ where + field + 𝑣 : ∀ k → {k ≤ n} → A + 𝑝 : ∀ k → ∀ {k≤n sk≤n} → 𝑣 k {k≤n} ≡ 𝑣 (suc k) {sk≤n} + + abstract + 𝑣₀ : A + 𝑣₀ = 𝑣 0 + + 𝑣₁ : {1 ≤ n} → A + 𝑣₁ {k≤} = 𝑣 1 {k≤} + + 𝑣₂ : {2 ≤ n} → A + 𝑣₂ {k≤} = 𝑣 2 {k≤} + + 𝑣₃ : {3 ≤ n} → A + 𝑣₃ {k≤} = 𝑣 3 {k≤} + + 𝑣₄ : {4 ≤ n} → A + 𝑣₄ {k≤} = 𝑣 4 {k≤} + + 𝑣₅ : {5 ≤ n} → A + 𝑣₅ {k≤} = 𝑣 5 {k≤} + + 𝑣₆ : {6 ≤ n} → A + 𝑣₆ {k≤} = 𝑣 6 {k≤} + + 𝑣₇ : {7 ≤ n} → A + 𝑣₇ {k≤} = 𝑣 7 {k≤} + + + 𝑝₀ : ∀ {sk≤n} → 𝑣₀ ≡ 𝑣₁ {sk≤n} + 𝑝₀ {sk≤n} = 𝑝 zero {tt} {sk≤n} + + 𝑝₁ : ∀ {k≤n sk≤n} → 𝑣₁ {k≤n} ≡ 𝑣₂ {sk≤n} + 𝑝₁ {k≤n} {sk≤n} = 𝑝 1 {k≤n} {sk≤n} + + 𝑝₂ : ∀ {k≤n sk≤n} → 𝑣₂ {k≤n} ≡ 𝑣₃ {sk≤n} + 𝑝₂ {k≤n} {sk≤n} = 𝑝 2 {k≤n} {sk≤n} + + 𝑝₃ : ∀ {k≤n sk≤n} → 𝑣₃ {k≤n} ≡ 𝑣₄ {sk≤n} + 𝑝₃ {k≤n} {sk≤n} = 𝑝 3 {k≤n} {sk≤n} + + 𝑝₄ : ∀ {k≤n sk≤n} → 𝑣₄ {k≤n} ≡ 𝑣₅ {sk≤n} + 𝑝₄ {k≤n} {sk≤n} = 𝑝 4 {k≤n} {sk≤n} + + 𝑝₅ : ∀ {k≤n sk≤n} → 𝑣₅ {k≤n} ≡ 𝑣₆ {sk≤n} + 𝑝₅ {k≤n} {sk≤n} = 𝑝 5 {k≤n} {sk≤n} + + 𝑝₆ : ∀ {k≤n sk≤n} → 𝑣₆ {k≤n} ≡ 𝑣₇ {sk≤n} + 𝑝₆ {k≤n} {sk≤n} = 𝑝 6 {k≤n} {sk≤n} + +data Sequence (n : ℕ) : Type where + 𝓿 : ∀ k → {k ≤ n} → Sequence n + 𝓹 : ∀ k → ∀ {k≤n sk≤n} → 𝓿 k {k≤n} ≡ 𝓿 (suc k) {sk≤n} + + +module _ {ℓ} n (A : Type ℓ) where + + fromNPath : (Sequence n → A) → NPath n A + fromNPath x .NPath.𝑣 k {k≤n} = x (𝓿 k {k≤n}) + fromNPath x .NPath.𝑝 k {k≤n} {k≤n'} i = x (𝓹 k {k≤n} {k≤n'} i) + + toNPath : NPath n A → (Sequence n → A) + toNPath x (𝓿 k {k≤n}) = x .NPath.𝑣 k {k≤n} + toNPath x (𝓹 k {k≤n} {k≤n'} i) = x .NPath.𝑝 k {k≤n} {k≤n'} i + + IsoFunSequenceNPath : Iso (NPath n A) (Sequence n → A) + IsoFunSequenceNPath .Iso.fun = toNPath + IsoFunSequenceNPath .Iso.inv = fromNPath + IsoFunSequenceNPath .Iso.rightInv b i a@(𝓿 k) = b a + IsoFunSequenceNPath .Iso.rightInv b i a@(𝓹 k i₁) = b a + IsoFunSequenceNPath .Iso.leftInv a i .NPath.𝑣 = a .NPath.𝑣 + IsoFunSequenceNPath .Iso.leftInv a i .NPath.𝑝 = a .NPath.𝑝 + + + + + + + +_∙f0_ : ∀ {ℓ} {A : Type ℓ} → + ∀ {x₀ y₀ y₁ z : A} + → {p₀ : x₀ ≡ y₀} {q₀ : y₀ ≡ z} {q₁ : y₁ ≡ z} + → {r : x₀ ≡ y₁} {y≡ : y₀ ≡ y₁} + → Square p₀ (λ _ → y₁) r y≡ + → Square q₀ q₁ y≡ (λ _ → z) + + → Square (p₀ ∙' q₀) q₁ r λ _ → z +(u ∙f0 v) j i = + hcomp (λ k → primPOr (~ i) (i ∨ j) (λ _ → u j (~ k)) λ _ → v j i) + (v j i) + +_∙f1_ : ∀ {ℓ} {A : Type ℓ} → + ∀ {x₁ y₀ y₁ z : A} + → {p₁ : x₁ ≡ y₁} {q₀ : y₀ ≡ z} {q₁ : y₁ ≡ z} + → {r : y₀ ≡ x₁} {y≡ : y₀ ≡ y₁} + → Square (λ _ → y₀) p₁ r y≡ + → Square q₀ q₁ y≡ (λ _ → z) + + → Square q₀ (p₁ ∙' q₁) r λ _ → z +(u ∙f1 v) j i = + hcomp (λ k → primPOr (~ i) (i ∨ (~ j)) (λ _ → u j (~ k)) λ _ → v j i) + (v j i) + + +_∙∙■_∙∙■_ : ∀ {ℓ} {A : Type ℓ} → + ∀ {x x₀ x₁ x₂ x₃ : A} + → {p₀ : x₀ ≡ x₁} {p₁ : x₁ ≡ x₂} {p₂ : x₂ ≡ x₃} + {q₀ : x₀ ≡ x} {q₁ : x₁ ≡ x} {q₂ : x₂ ≡ x} {q₃ : x₃ ≡ x} + → Square q₀ q₁ p₀ refl + → Square q₁ q₂ p₁ refl + → Square q₂ q₃ p₂ refl + → Square q₀ q₃ (p₀ ∙∙ p₁ ∙∙ p₂) refl +_∙∙■_∙∙■_ {x = x} s₀ s₁ s₂ j i = + hcomp (λ k → λ where + (j = i0) → s₀ (~ k) i + (j = i1) → s₂ k i + (i = i1) → x + ) (s₁ j i) + +◪→≡ : ∀ {ℓ} {A : Type ℓ} {x y : A} {p p' : x ≡ y} → + Square p' refl p refl → p ≡ p' +◪→≡ s i j = + hcomp (λ k → λ where + (j = i0) → s i0 (i ∧ ~ k) + (i = i1) → s i0 (j ∨ ~ k) + (i = i0) → s j i ; (j = i1) → s j i) (s j i) + +◪→◪→≡ : ∀ {ℓ} {A : Type ℓ} {x y : A} {p₀ p₁ p : x ≡ y} + → Square p refl p₀ refl + → Square p refl p₁ refl + → p₀ ≡ p₁ +◪→◪→≡ {y = y} {p = p} s₀ s₁ i j = + hcomp + (λ k → primPOr (~ i ∨ ~ j ∨ j) i (λ _ → s₀ j (~ k)) + λ _ → s₁ j (~ k)) y + +◪→◪→≡' : ∀ {ℓ} {A : Type ℓ} {x y : A} {p₀ p₁ p : x ≡ y} + → Square refl p refl p₀ + → Square refl p refl p₁ + → p₀ ≡ p₁ +◪→◪→≡' {y = y} {p = p} s₀ s₁ i j = + ◪→◪→≡ (λ i₁ i₂ → s₀ (~ i₁) (~ i₂)) + (λ i₁ i₂ → s₁ (~ i₁) (~ i₂)) i (~ j) + +comp₋₀ : ∀ {ℓ} {A : Type ℓ} → + {a a₀₀ : A} {a₀₋ : a₀₀ ≡ a} + {a₁₀ : A} {a₁₋ : a₁₀ ≡ a} + {a₋₀ a₋₀' : a₀₀ ≡ a₁₀} + → Square a₀₋ a₁₋ a₋₀ refl + → a₋₀' ≡ a₋₀ + → Square a₀₋ a₁₋ a₋₀' refl +comp₋₀ s p i j = + hcomp (λ k → primPOr (~ j) (j ∨ i ∨ ~ i) (λ _ → p (~ k) i) λ _ → s i j) (s i j) + +◪mkSq : ∀ {ℓ} {A : Type ℓ} → + {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} + {a₁₀ a₁₁ : A} {a₁₋ p₁₀ : a₁₀ ≡ a₁₁} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ p₀₁ : a₀₁ ≡ a₁₁} + → {p : a₀₀ ≡ a₁₁} + → Square p p₀₁ a₀₋ refl + → Square p₁₀ refl a₁₋ refl + → Square p p₁₀ a₋₀ refl + → Square p₀₁ refl a₋₁ refl + → Square a₀₋ a₁₋ a₋₀ a₋₁ +◪mkSq {a₁₁ = a₁₁} s₀₋ s₁₋ s₋₀ s₋₁ i j = + hcomp (λ k → λ where + (i = i0) → s₀₋ j (~ k) + (i = i1) → s₁₋ j (~ k) + (j = i0) → s₋₀ i (~ k) + (j = i1) → s₋₁ i (~ k)) + a₁₁ + + +coh₃helperI : ∀ {ℓ} {A : Type ℓ} → + {x₀ x₁ : A} → {p p₀₀ p₀₁ p₁₀ p₁₁ : x₀ ≡ x₁} → + {s₀₀ : Square refl p₀₀ refl p} + {s₀₁ : Square refl p₀₁ refl p} + {s₁₀ : Square refl p₁₀ refl p} + {s₁₁ : Square refl p₁₁ refl p} + → I → I → I → A +coh₃helperI {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} i j k = + hcomp + (λ z → λ { + (k = i0) → x₀ + ;(k = i1) → p z + ;(i = i0)(j = i0) → s₀₀ z k + ;(i = i1)(j = i0) → s₁₀ z k + ;(i = i0)(j = i1) → s₀₁ z k + ;(i = i1)(j = i1) → s₁₁ z k + }) x₀ + +coh₃helper : ∀ {ℓ} {A : Type ℓ} → + {x₀ x₁ : A} → {p p₀₀ p₀₁ p₁₀ p₁₁ : x₀ ≡ x₁} → + {s₀₀ : Square refl p₀₀ refl p} + {s₀₁ : Square refl p₀₁ refl p} + {s₁₀ : Square refl p₁₀ refl p} + {s₁₁ : Square refl p₁₁ refl p} + → + Cube + (λ j k → coh₃helperI {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} i0 j k) + (λ j k → coh₃helperI {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} i1 j k) + (λ j k → coh₃helperI {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} j i0 k) + (λ j k → coh₃helperI {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} j i1 k) + (λ _ _ → x₀) (λ _ _ → x₁) +coh₃helper {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} i j k = + coh₃helperI + {x₀ = x₀} {p = p} {s₀₀ = s₀₀} {s₀₁} {s₁₀} {s₁₁} i j k + +comp-coh-helper : ∀ {ℓ} {A : Type ℓ} → + {x₀ x₁ : A} → {p p₀ p₁ p₂ : x₀ ≡ x₁} → + {s₀ : Square refl p₀ refl p} + {s₁ : Square refl p₁ refl p} + {s₂ : Square refl p₂ refl p} + → Cube _ _ _ _ _ _ +comp-coh-helper {x₀ = x₀} {x₁} {p = p} {p₀} {p₁} {p₂} {s₀ = s₀} {s₁} {s₂} = + λ k i j → + hcomp + (λ z → λ { + (j = i0) → x₀ + ;(j = i1) → p (~ k ∨ z ∨ ~ i) + ;(i = i0) → p₀ j + ;(i = i1) → hcomp (λ k' → λ {(z = i0) → s₁ (k' ∧ ~ k) j + ;(z = i1) → s₂ k' j + ;(j = i0) → x₀ + ;(j = i1) → p (k' ∧ (~ k ∨ z)) + }) x₀ + ;(k = i1) → hcomp + (λ k → λ {(i = i0) → s₀ k j + ;(i = i1)(z = i0) → x₀ + ;(i = i1)(z = i1) → s₂ k j + ;(j = i0) → x₀ + ;(j = i1) → p (k ∧ (z ∨ ~ i)) + }) x₀ + + }) (hcomp (λ k' → λ {(i = i0) → s₀ k' j + ;(i = i1) → s₁ (k' ∧ ~ k) j + ;(j = i0) → x₀ + ;(j = i1) → p (k' ∧ (~ k ∨ ~ i)) + }) x₀) + + +coh-cub : ∀ {ℓ} {A : Type ℓ} {a₀₀₀ a₀₀₁ : A} {a₀₀₋ p₀₀₁ : a₀₀₀ ≡ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ p₀₁₋' : a₀₁₀ ≡ a₀₁₁} + {a₀₋₀ p₀₁₀ : a₀₀₀ ≡ a₀₁₀} {a₀₋₁ p₀₋₁' : a₀₀₁ ≡ a₀₁₁} + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ p₁₀₋' : a₁₀₀ ≡ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ ≡ a₁₁₁} + {a₁₋₀ p₁₋₀' : a₁₀₀ ≡ a₁₁₀} {a₁₋₁ : a₁₀₁ ≡ a₁₁₁} + {a₋₀₀ p₁₀₀ : a₀₀₀ ≡ a₁₀₀} {a₋₀₁ p₋₀₁' : a₀₀₁ ≡ a₁₀₁} + {a₋₁₀ p₋₁₀' : a₀₁₀ ≡ a₁₁₀} {a₋₁₁ : a₀₁₁ ≡ a₁₁₁} + + {p₀₁₁ : a₀₀₀ ≡ a₀₁₁} + {p₁₀₁ : a₀₀₀ ≡ a₁₀₁} + {p₁₁₀ : a₀₀₀ ≡ a₁₁₀} + {s₋₀₀ : Square refl p₁₀₀ refl a₋₀₀} + {s₋₀₁ : Square p₀₀₁ p₁₀₁ refl a₋₀₁} + {s₋₁₀ : Square p₀₁₀ p₁₁₀ refl a₋₁₀} + + {s₀₀₋ : Square refl p₀₀₁ refl a₀₀₋} + {s₀₁₋ : Square p₀₁₀ p₀₁₁ refl a₀₁₋} + {s₁₀₋ : Square p₁₀₀ p₁₀₁ refl a₁₀₋} + + {s₀₋₀ : Square refl p₀₁₀ refl a₀₋₀} + {s₀₋₁ : Square p₀₀₁ p₀₁₁ refl a₀₋₁} + {s₁₋₀ : Square p₁₀₀ p₁₁₀ refl a₁₋₀} + + + {p₋₁₁ : a₁₀₀ ≡ a₁₁₁} + {s'₁₋₀ : Square refl p₁₋₀' refl a₁₋₀} + {s'₁₀₋ : Square refl p₁₀₋' refl a₁₀₋} + {s'ₓ₁₋ : Square p₁₋₀' p₋₁₁ refl a₁₁₋} + {s'ₓ₋₁ : Square p₁₀₋' p₋₁₁ refl a₁₋₁} + + + {p₁₋₁ : a₀₁₀ ≡ a₁₁₁} + {s'₀₁₋ : Square refl p₀₁₋' refl a₀₁₋} + {s'₋₁₀ : Square refl p₋₁₀' refl a₋₁₀} + {s'₁ₓ₋ : Square p₋₁₀' p₁₋₁ refl a₁₁₋} + {s'₋ₓ₁ : Square p₀₁₋' p₁₋₁ refl a₋₁₁} + + + {p₁₁₋ : a₀₀₁ ≡ a₁₁₁} + {s'₋₀₁ : Square refl p₋₀₁' refl a₋₀₁} + {s'₀₋₁ : Square refl p₀₋₁' refl a₀₋₁} + {s'₁₋ₓ : Square p₋₀₁' p₁₁₋ refl a₁₋₁} + {s'₋₁ₓ : Square p₀₋₁' p₁₁₋ refl a₋₁₁} + + (s₀q₁ : _) + (c₀u₁ : Cube + (λ z z' → p₁₀₀ (~ z' ∨ z )) s₀q₁ + (λ _ z' → p₁₀₀ (~ z')) s'₁₀₋ + (λ _ _ → a₁₀₀) s₁₀₋ ) + (s₀q₂ : _) + (c₀u₂ : Cube + (λ z z' → p₁₀₀ (~ z' ∨ z )) s₀q₂ + (λ _ z' → p₁₀₀ (~ z')) s'₁₋₀ + (λ _ _ → a₁₀₀) s₁₋₀ ) + + (s₁q₀ : _) + (c₁u₀ : Cube + (λ z z' → p₀₁₀ (~ z' ∨ z )) s₁q₀ + (λ _ z' → p₀₁₀ (~ z')) s'₀₁₋ + (λ _ _ → a₀₁₀) s₀₁₋ ) + (s₁q₂ : _) + (c₁u₂ : Cube + (λ z z' → p₀₁₀ (~ z' ∨ z )) s₁q₂ + (λ _ z' → p₀₁₀ (~ z')) s'₋₁₀ + (λ _ _ → a₀₁₀) s₋₁₀ ) + + (s₂q₀ : _) + (c₂u₀ : Cube + (λ z z' → p₀₀₁ (~ z' ∨ z )) s₂q₀ + (λ _ z' → p₀₀₁ (~ z')) s'₀₋₁ + (λ _ _ → a₀₀₁) s₀₋₁ ) + (s₂q₁ : _) + (c₂u₁ : Cube + (λ z z' → p₀₀₁ (~ z' ∨ z )) s₂q₁ + (λ _ z' → p₀₀₁ (~ z')) s'₋₀₁ + (λ _ _ → a₀₀₁) s₋₀₁ ) + {p₁₁₁ : a₀₀₀ ≡ a₁₁₁} + (sp₀ : Square (sym p₁₀₀) p₋₁₁ refl p₁₁₁) + (sp₁ : Square (sym p₀₁₀) p₁₋₁ refl p₁₁₁) + (sp₂ : Square (sym p₀₀₁) p₁₁₋ refl p₁₁₁) + (sp₀₁ : Square p₁₁₀ p₁₁₁ refl a₁₁₋) + (sp₀₂ : Square p₁₀₁ p₁₁₁ refl a₁₋₁) + (sp₁₂ : Square p₀₁₁ p₁₁₁ refl a₋₁₁) + (s₀s₁ : Cube + s₀q₂ sp₀ + (λ _ z' → p₁₀₀ (~ z')) s'ₓ₁₋ + (λ _ _ → a₁₀₀) sp₀₁) + (s₀s₂ : Cube + s₀q₁ sp₀ + (λ _ z' → p₁₀₀ (~ z')) s'ₓ₋₁ + (λ _ _ → a₁₀₀) sp₀₂) + + (s₁s₀ : Cube + s₁q₂ sp₁ + (λ _ z' → p₀₁₀ (~ z')) s'₁ₓ₋ + (λ _ _ → a₀₁₀) sp₀₁) + (s₁s₂ : Cube + s₁q₀ sp₁ + (λ _ z' → p₀₁₀ (~ z')) s'₋ₓ₁ + (λ _ _ → a₀₁₀) sp₁₂) + + (s₂s₀ : Cube + s₂q₁ sp₂ + (λ _ z' → p₀₀₁ (~ z')) s'₁₋ₓ + (λ _ _ → a₀₀₁) sp₀₂) + (s₂s₁ : Cube + s₂q₀ sp₂ + (λ _ z' → p₀₀₁ (~ z')) s'₋₁ₓ + (λ _ _ → a₀₀₁) sp₁₂) + + → Cube + + (λ i₁ i₂ → ◪mkSq (λ u v → s₀₁₋ (~ u) (~ v)) (λ u v → s₀₀₋ (~ u) (~ v)) + (λ u v → s₀₋₁ (~ u) (~ v)) (λ u v → s₀₋₀ (~ u) (~ v)) (~ i₁) (~ i₂)) + + (λ i₁ i₂ → ◪mkSq (λ u v → s'ₓ₁₋ (~ u) (~ v)) (λ u v → s'₁₀₋ (~ u) (~ v)) + (λ u v → s'ₓ₋₁ (~ u) (~ v)) ((λ u v → s'₁₋₀ (~ u) (~ v))) (~ i₁) (~ i₂)) + + (λ i₀ i₂ → ◪mkSq (λ u v → s₁₀₋ (~ u) (~ v)) ((λ u v → s₀₀₋ (~ u) (~ v))) + ((λ u v → s₋₀₁ (~ u) (~ v))) ((λ u v → s₋₀₀ (~ u) (~ v))) (~ i₀) (~ i₂)) + + + + (λ i₀ i₂ → ◪mkSq (λ u v → s'₁ₓ₋ (~ u) (~ v)) (λ u v → s'₀₁₋ (~ u) (~ v)) + (λ u v → s'₋ₓ₁ (~ u) (~ v)) (λ u v → s'₋₁₀ (~ u) (~ v)) (~ i₀) (~ i₂)) + (λ i₀ i₁ → ◪mkSq ((λ u v → s₁₋₀ (~ u) (~ v))) (((λ u v → s₀₋₀ (~ u) (~ v)))) + ((λ u v → s₋₁₀ (~ u) (~ v))) (((λ u v → s₋₀₀ (~ u) (~ v)))) (~ i₀) (~ i₁)) + (λ i₀ i₁ → ◪mkSq (λ u v → s'₁₋ₓ (~ u) (~ v)) (λ u v → s'₀₋₁ (~ u) (~ v)) + (λ u v → s'₋₁ₓ (~ u) (~ v)) (λ u v → s'₋₀₁ (~ u) (~ v)) (~ i₀) (~ i₁)) +coh-cub {a₀₀₀ = a₀₀₀} {p₀₀₁ = p₀₀₁} {p₀₁₀ = p₀₁₀} {p₁₀₀ = p₁₀₀} {s₋₀₀ = s₋₀₀} {s₀₀₋ = s₀₀₋} {s₀₋₀ = s₀₋₀} + s₀q₁ c₀u₁ s₀q₂ c₀u₂ + s₁q₀ c₁u₀ s₁q₂ c₁u₂ + s₂q₀ c₂u₀ s₂q₁ c₂u₁ + sp₀ sp₁ sp₂ + sp₀₁ sp₀₂ sp₁₂ + s₀s₁ s₀s₂ + s₁s₀ s₁s₂ + s₂s₀ s₂s₁ + i₀ i₁ i₂ = + hcomp + (λ z → λ { + (i₁ = i0)(i₂ = i0) → s₋₀₀ i₀ z + ;(i₀ = i0)(i₂ = i0) → s₀₋₀ i₁ z + ;(i₀ = i0)(i₁ = i0) → s₀₀₋ i₂ z + ;(i₀ = i1) → + hcomp + (λ z' → λ { + (i₁ = i0) → c₀u₁ i₂ z z' + ;(i₂ = i0) → c₀u₂ i₁ z z' + ;(i₁ = i1) → s₀s₁ i₂ z z' + ;(i₂ = i1) → s₀s₂ i₁ z z' + ;(z = i0) → p₁₀₀ (~ z') + }) + (p₁₀₀ i1) + ;(i₁ = i1) → hcomp + (λ z' → λ { + (i₀ = i0) → c₁u₀ i₂ z z' + ;(i₂ = i0) → c₁u₂ i₀ z z' + ;(i₀ = i1) → s₁s₀ i₂ z z' + ;(i₂ = i1) → s₁s₂ i₀ z z' + ;(z = i0) → p₀₁₀ (~ z') + }) + (p₀₁₀ i1) + ;(i₂ = i1) → hcomp + (λ z' → λ { + (i₀ = i0) → c₂u₀ i₁ z z' + ;(i₁ = i0) → c₂u₁ i₀ z z' + ;(i₀ = i1) → s₂s₀ i₁ z z' + ;(i₁ = i1) → s₂s₁ i₀ z z' + ;(z = i0) → p₀₀₁ (~ z') + }) + (p₀₀₁ i1) + }) + a₀₀₀ diff --git a/Cubical/Tactics/PathSolver/Reflection.agda b/Cubical/Tactics/PathSolver/Reflection.agda new file mode 100644 index 0000000000..381489d866 --- /dev/null +++ b/Cubical/Tactics/PathSolver/Reflection.agda @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} + +module Cubical.Tactics.PathSolver.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 as L +open import Cubical.Data.Maybe as Mb + +open import Cubical.Reflection.Base +import Agda.Builtin.Reflection as R +open import Cubical.Tactics.Reflection + + + +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∙' x y = R.def (quote _∙'_) (x v∷ y v∷ [] ) diff --git a/Cubical/Tactics/Reflection.agda b/Cubical/Tactics/Reflection.agda index 9d48933be4..eb446a988f 100644 --- a/Cubical/Tactics/Reflection.agda +++ b/Cubical/Tactics/Reflection.agda @@ -9,6 +9,7 @@ module Cubical.Tactics.Reflection where open import Cubical.Foundations.Prelude open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String open import Cubical.Data.Bool open import Cubical.Data.List @@ -22,18 +23,52 @@ private variable ℓ ℓ' : Level -_<$>_ : ∀ {ℓ ℓ'} {A : Type ℓ}{B : Type ℓ'} → (A → B) → TC A → TC B -f <$> t = t >>= λ x → returnTC (f x) +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 -_<*>_ : ∀ {ℓ ℓ'} {A : Type ℓ}{B : Type ℓ'} → TC (A → B) → TC A → TC B -s <*> t = s >>= λ f → t >>= λ x → returnTC (f x) -wait-for-args : List (Arg Term) → TC (List (Arg Term)) +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 @@ -45,9 +80,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 @@ -85,6 +120,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 ( @@ -111,3 +152,9 @@ equation-solver don't-Reduce mk-call debug hole = unify-with-goal <|> typeError ((strErr "Could not equate the following expressions:\n " ∷ termErr elhs ∷ strErr "\nAnd\n " ∷ termErr erhs ∷ [])))))) + +unifyFromString : String → Term → TC Unit +unifyFromString s h = do + ty ← inferType h + t ← checkFromStringTC s ty + unify t h diff --git a/Cubical/Tactics/Reflection/CuTerm.agda b/Cubical/Tactics/Reflection/CuTerm.agda new file mode 100644 index 0000000000..a8e2318ad2 --- /dev/null +++ b/Cubical/Tactics/Reflection/CuTerm.agda @@ -0,0 +1,591 @@ +{- + +This module provides a representation of reflected syntax focused on terms involving compositions. + +The representation can be thought as version of `Term` with additional nodes, capturing the composition operations and partial elements. + +Definition is parameterized by two types: +* `CongGuard`: A type that can be either `Unit` or `⊥`, acting as a flag to enable or disable specific operations over composition terms. +* `A`: A type that allows attaching specific data to the faces of the cubical complex, represented by nested compositions. + +Main definitions and functions: + +* `CuTerm'`: The core datatype representing cubical terms, parameterized by `CongGuard` and `A`. + - `hco'`: Constructor for hcomp node. + - `cell'`: Constructor for a `leaf` cell - `Term` with associated type `A`. + - `𝒄ong'`: Constructor for terms which are not `hcomp` , but contains `hcomp` as some subterms. + This is the constructor which is guarded by `CongGuard` + +* `CuTerm` and `CuTermNC`: Specializations of `CuTerm'` with specific `CongGuard` values. + +* Utility functions: + - `isCell`: Checks if a `CuTerm` is a cell. + - `is𝑪ongFree`, `is𝑪ongFreeF`: Checks if a `CuTerm` or a list of terms are free of `𝒄ong`. + - `cellQ`: Checks if a `CuTerm` is a cell. + - `almostLeafQ`: Checks if a `CuTerm` is an almost leaf node in its structure. + +* Modules and utilities for rendering and term conversion: + - `prettyPrinter`: Pretty prints `CuTerm'` for diagnostic purposes. + - `ToTerm`: Converts `CuTerm'` into `R.Term` for further processing or evaluation. + +* Normalization and evaluation: + - `cuEval`, `cuEvalL`: Evaluates a `CuTerm'` given a subface and specific parameters. + - `normaliseCells`: Normalizes cells within `CuTerm'`. + +* Miscellaneous functions: + - `pickSFfromPartial'`, `pickSFfromPartial`: Picks a subface from a list given a partial value. + - `permuteVars`: Permutes variables within a `CuTerm` by a given function. + - `foldCells`, `visitCellsM`: Utility functions to traverse and apply transformations or checks over the cells. + - `tryCastAsNoCong`, `tryCastAsNoCongS`: Attempts to cast `CuTerm` to form where hcomps apears only ont the `top` level + +* Code generation: + - `codeGen`: Generates code from `CuTerm'` - this is version of pretty printer producing valid agda code, suitable for use within `checkFromStringTC`. + +Special annotations and constructs like `MetaTag` and `⁇` aid in reflection-based manipulation of terms and help manage intermediate or unspecified terms within the system. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.Reflection.CuTerm where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Data.List as L +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Bool +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.Nat +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) + + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.Error + +open import Cubical.Tactics.Reflection.Variables +open import Cubical.Tactics.Reflection.Utilities + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar + +private + variable + ℓ : Level + A B : Type ℓ + +data CuTerm' (CongGuard : Type) (A : Type ℓ) : Type ℓ + + +record Hco (CongGuard : Type) (A : Type ℓ) : Type ℓ where + inductive + constructor hcodata + field + sides : List (SubFace × CuTerm' CongGuard A) + bottom : CuTerm' CongGuard A + +data CuTerm' CongGuard A where + hco' : Hco CongGuard A → CuTerm' CongGuard A + cell' : A → R.Term → CuTerm' CongGuard A + 𝒄ong' : {cg : CongGuard} → R.Term → List ((Hco CongGuard A)) → CuTerm' CongGuard A + +pattern + cell x = cell' tt x + +pattern + hco x y = hco' (hcodata x y) + +pattern + 𝒄ong th tl = 𝒄ong' {cg = tt} th tl + +𝒄ongF : ∀ {CongGuard} {A : Type ℓ} {cg : CongGuard} {a : A} → R.Term → List ((CuTerm' CongGuard A)) → CuTerm' CongGuard A +𝒄ongF {cg = cg} {a = a} t xs = uncurry h (snd (foldl g (length xs , (t , [])) xs)) + + where + h : R.Term → List (Hco _ _) → CuTerm' _ _ + h t [] = cell' a t + h = 𝒄ong' {cg = cg} + + g : ℕ × (R.Term × List (Hco _ _)) → CuTerm' _ _ → ℕ × (R.Term × (List (Hco _ _))) + g (n , (t , xs)) (hco' x) = n , (t , xs ∷ʳ x) + g (n , (t , xs)) (cell' x x₁) = + predℕ n , replaceAtTrm (length xs) (liftVarsFrom n zero x₁) t , xs + g (n , (t , xs)) (𝒄ong' t' xs') = + (predℕ n + length xs') , + replaceAtTrm (length xs) + (liftVarsFrom (n ∸ suc (length xs)) ((length xs') + suc (length xs)) + $ liftVarsFrom (suc (length xs)) zero t') + (liftVarsFrom (length xs') (suc (length xs)) t) , + xs ++ xs' + +CuTerm = CuTerm' Unit Unit + +CuTermNC = CuTerm' ⊥ Unit + + +HcoNC→Hco : List (SubFace × CuTerm' ⊥ Unit) → List (SubFace × CuTerm' Unit Unit) + +CuTermNC→CuTerm : CuTermNC → CuTerm +CuTermNC→CuTerm (hco sides bottom) = hco (HcoNC→Hco sides) (CuTermNC→CuTerm bottom) +CuTermNC→CuTerm (cell' x x₁) = cell' x x₁ + +HcoNC→Hco [] = [] +HcoNC→Hco ((sf , x) ∷ xs) = (sf , CuTermNC→CuTerm x) ∷ HcoNC→Hco xs + + +isCell : CuTerm → Bool +isCell (cell x) = true +isCell _ = false + + +is𝑪ongFreeF : List (SubFace × CuTerm) → Bool + +is𝑪ongFree : CuTerm → Bool +is𝑪ongFree (hco x x₁) = is𝑪ongFreeF x and is𝑪ongFree x₁ +is𝑪ongFree (cell x) = true +is𝑪ongFree (𝒄ong x x₁) = false + +is𝑪ongFreeF [] = true +is𝑪ongFreeF ((_ , x) ∷ xs) = is𝑪ongFree x and is𝑪ongFreeF xs + + +cellQ : CuTerm → Bool +cellQ (cell x) = true +cellQ _ = false + +almostLeafQ : CuTerm → Bool +almostLeafQ (hco x (hco x₁ x₂)) = false +almostLeafQ (hco x (cell x₁)) = + foldr (_and_ ∘S cellQ ∘S snd) true x +almostLeafQ _ = false + + + +module prettyPrinter {A B : Type} (cellTermRender : CuCtx → R.Term → R.TC (List R.ErrorPart)) (dim : ℕ) where + + renderSubFaceExp : SubFace → R.TC String + renderSubFaceExp sf = R.normalise (SubFace→Term sf) >>= renderTerm + + + renderSubFacePattern : CuCtx → SubFace → String + renderSubFacePattern ctx sf = + foldl _<>_ "" (L.map + ((λ (b , k) → let k' = L.lookupAlways "‼" + (freeVars ctx) k + in "(" <> k' <> "=" <> (if b then "1" else "0") <> ")")) + (subFaceConstraints sf)) + -- (mapM (λ (b , k) → do k' ← renderTerm (R.var k []) + -- pure $ "(" <> k' <> "=" <> (if b then "1" else "0") <> ")") + -- ∘S subFaceConstraints) >=& foldl _<>_ "" + + ppCT'' : CuCtx → ℕ → CuTerm' A B → R.TC (List R.ErrorPart) + -- ppCArg : CuCtx → ℕ → CuArg → R.TC (List R.ErrorPart) + + ppCT'' _ zero _ = R.typeError [ "pPCT FAIL" ]ₑ + ppCT'' ctx (suc d) (hco x x₁) = do + let l = length ctx ∸ dim + indN ← foldr max zero <$> ( + (mapM ((((pure ∘ (renderSubFacePattern ctx)) >=& stringLength)) ∘S fst ) x)) + + let newDimVar = (mkNiceVar' "𝒛" l) + rest ← (L.intersperse (R.strErr "\n") ∘S L.join) <$> mapM + (λ (sf , cu) → do + + ( do + let sfTm = renderSubFacePattern ctx sf + (do sfTm' ← inCuCtx' (("z" , nothing) ∷ ctx) $ R.formatErrorParts [ liftVars (SubFace→TermInCtx ctx sf) ]ₑ + cu' ← (ppCT'' ((newDimVar , nothing) ∷ applyFaceConstraints sf ctx) d cu) + cu'' ← R.formatErrorParts cu' + let cu''' = indent' false ' ' 2 cu'' + pure (offsetStrR indN sfTm ∷ₑ + -- "/" ∷ₑ sfTm' ∷ₑ + " → " ∷ₑ [ cu''' ]ₑ))) >>= + (R.formatErrorParts >=& [_]ₑ)) x + lid ← indent '│' 1 <$> (ppCT'' ctx d x₁ >>= R.formatErrorParts) + rest' ← indent '║' 2 <$> R.formatErrorParts rest + pure $ (R.strErr ("\n𝒉𝒄𝒐𝒎𝒑 λ " <> newDimVar <> "\n")) ∷ + (rest' ∷ₑ "\n├─────────── \n" ∷ₑ + lid ∷ₑ [ "\n└─────────── "]ₑ) + + ppCT'' ctx _ (cell' _ x) = do + ctr ← cellTermRender ctx x >>= + --inCuCtx ctx ∘ + R.formatErrorParts + pure [ ctr ]ₑ + ppCT'' ctx (suc d) (𝒄ong' h t) = do + rT ← mapM ((argRndr ∘S hco') >=> R.formatErrorParts) t + rHead ← inCuCtx ctx $ addNDimsToCtx' "𝒙" (length t) $ renderTerm h + pure $ [ rHead <> indent ' ' 2 (foldr (_<>_ ∘S ("\n" <>_)) "" rT)]ₑ + + where + argRndr : CuTerm' A B → R.TC _ + argRndr x = + ((λ s → [ "(" ]ₑ ++ s ++ [ ")" ]ₑ) <$> (ppCT'' ctx d x)) + + ppCT' : ℕ → CuTerm' A B → R.TC (List R.ErrorPart) + ppCT' = ppCT'' (defaultCtx dim) + + +ppCTn : {A B : Type} → Bool → ℕ → ℕ → CuTerm' A B → R.TC (List R.ErrorPart) +ppCTn b = + prettyPrinter.ppCT' (λ ctx x → + do inCuCtx ctx $ do + nt ← (if b then R.normalise else R.reduce) x + x'' ← R.formatErrorParts [ nt ]ₑ + pure [ R.strErr x'' ]) + + +ppCT : {A B : Type} → ℕ → ℕ → CuTerm' A B → R.TC (List R.ErrorPart) +ppCT = ppCTn true + + +ppCTs : {A B : Type} → ℕ → ℕ → CuTerm' A B → R.TC (List R.ErrorPart) +ppCTs = prettyPrinter.ppCT' (λ _ x → pure [ R.strErr "■" ]) + + + +constPartial : A → ∀ φ → Partial φ A +constPartial a φ 1=1 = a + +module ToTerm {A B : Type} where + + toTerm : CuCtx → CuTerm' A B → R.Term + toTermFill toTermFill' : CuCtx → Hco A B → R.Term + + + toTermA : CuCtx → List (Hco A B) → List (R.Term) + + + mkSFTrm : CuCtx → SubFace × CuTerm' A B → R.Term + mkSFTrm ctx (sf , cu) = + R.def (quote constPartial) + (toTerm (("𝒛" , nothing) ∷ applyFaceConstraints sf ctx) cu v∷ + v[ liftVars (SubFace→TermInCtx ctx sf) ]) + + toSides : CuCtx → List (SubFace × CuTerm' A B) → R.Term + toSides ctx [] = R.pat-lam [] [] + toSides ctx (x ∷ []) = mkSFTrm ctx x + + + + toSides ctx ((sf , cu) ∷ xs@(_ ∷ _)) = + R.def (quote primPOr) + (liftVars (SubFace→TermInCtx ctx sf) v∷ R.unknown v∷ + (mkSFTrm ctx (sf , cu)) v∷ v[ toSides ctx xs ]) + + toTermA ctx [] = [] + toTermA ctx (x ∷ xs) = + (toTerm ctx (hco' x)) ∷ toTermA ctx xs + + toTerm ctx (hco' (hcodata x x₁)) = + R.def (quote hcomp) + (vlam "𝒛" (toSides ctx x) v∷ v[ toTerm ctx x₁ ]) + toTerm ctx (cell' _ x) = + liftWhere (L.map ((λ { (just _) → true ; _ → false }) ∘S snd ) ctx) x + + toTerm ctx (𝒄ong' h t) = + let h' = liftWhere (repeat (length t) false ++ L.map ((λ { (just _) → true ; _ → false }) ∘S snd ) ctx) h + in substTms (toTermA ctx (t)) h' + + toTermFill ctx (hcodata x x₁) = + R.def (quote hfill) + (liftVars (vlam "𝒛" (toSides ctx x)) v∷ + R.def (quote inS) v[ liftVars (toTerm ctx x₁) ] v∷ v[ 𝒗 zero ]) + + toTermFill' ctx (hcodata x x₁) = + R.def (quote hfill) + (liftVarsFrom 1 (length ctx) (vlam "𝒛" (toSides ctx x)) v∷ + R.def (quote inS) v[ liftVarsFrom 1 (length ctx) (toTerm ctx x₁) ] v∷ v[ 𝒗 (length ctx) ]) + +toTerm : {A B : Type} → ℕ → CuTerm' A B → R.Term +toTerm dim = vlamⁿ dim ∘ (ToTerm.toTerm (defaultCtx dim)) + + + +module cuEval {A : Type} {b : B} where + + cuEval : ℕ → SubFace → CuTerm' A B → CuTerm' A B + cuEvalL : ℕ → SubFace → List (Hco A B) → List (CuTerm' A B) + + cuEvalL _ sf [] = [] + cuEvalL fuel sf (x ∷ l) = cuEval fuel sf (hco' x) ∷ cuEvalL fuel sf l + cuEval zero _ _ = cell' b (R.lit (R.string "out of fuel in cuEval")) + cuEval (suc fuel) sf (hco l y) = + let sSf = findBy (⊂⊃? ∘S (sf _) ∘S fst) l + in h sSf + + where + msf : SubFace × CuTerm' A B → Maybe (SubFace × CuTerm' A B) + msf (sf' , t) = + map-Maybe + (λ sf'' → (injSF sf sf'') , cuEval fuel (nothing ∷ (injSF sf' sf'')) t) + (sf' sf∩ sf) + + h : Maybe (SubFace × CuTerm' A B) → CuTerm' A B + h (just (_ , x)) = cuEval fuel (just true ∷ repeat (sfDim sf) nothing) x + h nothing = + Mb.rec + (let l' = L.catMaybes (L.map msf l) + in hco l' (cuEval fuel sf y)) + (λ (sf' , f) → cuEval fuel (just true ∷ (injSF sf' sf)) f) + (findBy (sf ⊂?_ ∘S [_] ∘S fst) l) + + + + cuEval fuel sf (cell' x x₁) = cell' x (subfaceCell sf x₁) + cuEval fuel sf (𝒄ong' {cg = cg} h tl) = + let h' = subfaceCell (repeat (length tl) nothing ++ sf) h + in 𝒄ongF {cg = cg} {a = b} h' (cuEvalL fuel sf tl) + +cuEval : {A : Type} {B : Type ℓ} {b : B} → SubFace → CuTerm' A B → CuTerm' A B +cuEval {b = b} = cuEval.cuEval {b = b} 100 + + + +pickSFfromPartial' : A → SubFace → List (SubFace × CuTerm' B A) → Maybe (CuTerm' B A) +pickSFfromPartial' a sf l = + let sSf = findBy (sf ⊂?_ ∘S [_] ∘S fst) l + in map-Maybe (λ (sf' , f) → cuEval {b = a} (nothing ∷ (injSF sf' sf)) f) sSf + +pickSFfromPartial : SubFace → List (SubFace × CuTerm) → Maybe (CuTerm) +pickSFfromPartial = pickSFfromPartial' _ + + + + + +module normaliseCells where + + + ncH : ℕ → ℕ → (Hco A B) → R.TC (Hco A B) + + nc : ℕ → ℕ → (CuTerm' A B) → R.TC (CuTerm' A B) + nc zero _ _ = R.typeError [ "out of fuel in normaliceCells" ]ₑ + nc (suc fuel) dim (hco' x) = ⦇ hco' (ncH (fuel) dim x) ⦈ + + nc (suc fuel) dim (cell' x x₁) = + cell' x <$> (addNDimsToCtx dim $ R.normalise x₁) + nc (suc fuel) dim (𝒄ong' {cg = cg} x x₁) = + 𝒄ong' {cg = cg} x <$> mapM (ncH fuel dim) x₁ + + ncH fuel dim (hcodata x x₁) = + ⦇ hcodata + (mapM (λ (sf , x) → ⦇ ⦇ sf ⦈ , ( nc fuel (suc (sfDim sf)) x) ⦈ ) x) + (nc (fuel) dim x₁) ⦈ + +normaliseCells : ℕ → CuTerm' A B → R.TC (CuTerm' _ _) +normaliseCells = normaliseCells.nc 100 + +cuEvalN : SubFace → (CuTerm' A Unit) → R.TC (CuTerm' A Unit) +cuEvalN sf = normaliseCells (sfDim sf) ∘S cuEval sf + + +mostNestedCap : CuTermNC → R.Term +mostNestedCap (hco x x₁) = mostNestedCap x₁ +mostNestedCap (cell' x x₁) = x₁ + + +-- this can be trusted, only if we sure that term already typechecks! + +allCellsConstant? : ℕ → CuTerm' A B → Bool +allCellsConstant? dim x = h dim x + where + h : ℕ → CuTerm' _ _ → Bool + hs : List (SubFace × CuTerm' _ _) → Bool + + h dim (hco x₁ x₂) = h dim x₂ and hs x₁ + + h dim (cell' x₁ x₂) = not (hasVarBy (_<ℕ dim) x₂) + h dim (𝒄ong' x₁ x₂) = false + + hs [] = true + hs ((sf , x) ∷ xs) = (h (suc (sfDim sf)) x) and hs xs + +module permuteVars where + + permute : (ℕ → ℕ) → SubFace → SubFace + permute f sf = foldr (λ k → replaceAt (f k) (lookupAlways nothing sf k)) sf (range (length sf)) + + intoFace fromFace : SubFace → ℕ → ℕ + intoFace [] k = k + intoFace (nothing ∷ sf) zero = zero + intoFace (nothing ∷ sf) (suc k) = suc (intoFace sf k) + intoFace (just x ∷ sf) k = intoFace sf (predℕ k) + fromFace [] k = k + fromFace (nothing ∷ xs) zero = zero + fromFace (nothing ∷ xs) (suc k) = suc (fromFace xs k) + fromFace (just x ∷ xs) k = suc (fromFace xs k) + + sfPerm : SubFace → (ℕ → ℕ) → (ℕ → ℕ) + sfPerm sf f k = + if k =ℕ sfDim sf + then k + else intoFace (permute f sf) (f (fromFace sf k)) + + nc : ℕ → (ℕ → ℕ) → ℕ → CuTerm → R.TC CuTerm + nc zero _ _ _ = R.typeError [ "out of fuel in permuteVars" ]ₑ + nc (suc fuel) prm dim (hco x x₁) = do + + ⦇ hco + (mapM (λ (sf , c) → ⦇ ⦇ (permute prm sf) ⦈ , (nc fuel (sfPerm sf prm) (suc (sfDim sf)) c) ⦈) x) + (nc (suc fuel) prm dim x₁) ⦈ + nc (suc fuel) prm _ (cell x) = + pure $ cell (remapVars prm x) + nc (suc fuel) _ _ (𝒄ong' x x₁) = R.typeError [ "TODO in permuteVars" ]ₑ + + +permuteVars = permuteVars.nc 100 + + +CuBoundary' : ∀ A B → Type ℓ +CuBoundary' A B = List (CuTerm' A B × CuTerm' A B) + +CuBoundary = CuBoundary' Unit Unit + + +tryCastAsNoCongS : (List (SubFace × CuTerm)) → Maybe (List (SubFace × CuTerm' ⊥ Unit)) + + +tryCastAsNoCong : CuTerm → Maybe (CuTerm' ⊥ Unit) +tryCastAsNoCong (hco x x₁) = + ⦇ hco (tryCastAsNoCongS x) (tryCastAsNoCong x₁) ⦈ +tryCastAsNoCong (cell x) = pure $ cell' _ x +tryCastAsNoCong (𝒄ong' x x₁) = nothing + + +tryCastAsNoCongS [] = ⦇ [] ⦈ +tryCastAsNoCongS ((sf , x) ∷ xs) = + ⦇ (⦇ ⦇ sf ⦈ , (tryCastAsNoCong x) ⦈) ∷ (tryCastAsNoCongS xs) ⦈ + + +foldCells : (A → B → B) → CuTerm' ⊥ A → B → B +foldCells {A = A} {B = B} f = fc + where + fcs : List (SubFace × CuTerm' ⊥ A) → B → B + + fc : CuTerm' ⊥ A → B → B + fc (hco x x₂) b = fc x₂ (fcs x b) + fc (cell' x x₂) b = f x b + + fcs [] b = b + fcs ((_ , x) ∷ x₂) b = fcs x₂ (fc x b) + + +visitCellsM : (A → R.TC Unit) → CuTerm' ⊥ A → R.TC Unit +visitCellsM {A = A} f = vc + where + + vcs : List (SubFace × CuTerm' ⊥ A) → R.TC Unit + + vc : CuTerm' ⊥ A → R.TC Unit + vc (hco x x₁) = vc x₁ >> vcs x >> pure _ + vc (cell' x x₁) = f x + + vcs [] = pure _ + vcs ((_ , x) ∷ xs) = vc x >> vcs xs + +data MetaTag : Type where + +-- metaCell : CuTerm +pattern metaCell = cell (R.def (quote MetaTag) []) + + +module codeGen {A B : Type} (normaliseCells : Bool) (dim : ℕ) where + + renderSubFaceExp : SubFace → R.TC String + renderSubFaceExp sf = R.normalise (SubFace→Term sf) >>= renderTerm + + + + max-𝒛-idx : CuCtx → ℕ + max-𝒛-idx = foldr ((max ∘S (λ { (just ("𝒛" , k )) → (suc k) ; _ → zero }) ∘S getSubscript) ∘S fst ) zero + + renderSubFacePattern : CuCtx → SubFace → String + renderSubFacePattern ctx sf = + foldl _<>_ "" (L.map + ((λ (b , k) → let k' = L.lookupAlways "‼" + (freeVars ctx) k + in "(" <> k' <> " = " <> (if b then "i1" else "i0") <> ")")) + (subFaceConstraints sf)) + + ppCT'' : CuCtx → ℕ → CuTerm' A B → R.TC (List R.ErrorPart) + -- ppCArg : CuCtx → ℕ → CuArg → R.TC (List R.ErrorPart) + + ppCT'' _ zero _ = R.typeError [ "pPCT FAIL" ]ₑ + ppCT'' ctx (suc d) (hco x x₁) = do + let l = max (length ctx ∸ dim) (max-𝒛-idx ctx) + indN ← foldr max zero <$> ( + (mapM ((((pure ∘ (renderSubFacePattern ctx)) >=& stringLength)) ∘S fst ) x)) + + let newDimVar = (mkNiceVar' "𝒛" l) + rest ← (mapAt (λ { (R.strErr s) → R.strErr $ (" " <> s) ; x → x} ) zero + ∘S L.intersperse (R.strErr "\n;") ∘S L.join) <$> mapM + (λ (sf , cu) → do + + + + -- R.extendContext "zz" (varg (R.def (quote I) [])) $ + ( do + let sfTm = renderSubFacePattern ctx sf + -- R.extendContext newDimVar (varg (R.def (quote I) [])) $ + (do sfTm' ← inCuCtx' (("z" , nothing) ∷ ctx) $ R.formatErrorParts [ liftVars (SubFace→TermInCtx ctx sf) ]ₑ + cu' ← (ppCT'' ((newDimVar , nothing) ∷ applyFaceConstraints sf ctx) d cu) + cu'' ← R.formatErrorParts cu' + let cu''' = indent' false ' ' 2 cu'' + pure (offsetStrR indN sfTm ∷ₑ + -- "/" ∷ₑ sfTm' ∷ₑ + " → " ∷ₑ [ cu''' ]ₑ))) >>= + (R.formatErrorParts >=& [_]ₑ)) x + lid ← (trimLeft ∘S indent ' ' 1) <$> (ppCT'' ctx d x₁ >>= R.formatErrorParts) + rest' ← indent ' ' 2 <$> R.formatErrorParts rest + pure $ (R.strErr ("\nhcomp (λ " <> newDimVar <> " → λ { \n")) ∷ + (rest' ∷ₑ "\n }) \n" ∷ₑ + "(" ∷ₑ lid ∷ₑ ")" ∷ₑ [ "\n "]ₑ) + + ppCT'' ctx _ (cell' _ (R.def (quote MetaTag) [])) = pure [ R.strErr "?" ] + ppCT'' ctx _ (cell' _ x) = do + ctr ← inCuCtx ctx $ do + nt ← (if normaliseCells then R.normalise else pure) x + x'' ← R.formatErrorParts [ nt ]ₑ + pure [ R.strErr (x'') ] + -- cellTermRender ctx x >>= + -- --inCuCtx ctx ∘ + -- R.formatErrorParts + pure ctr + ppCT'' ctx (suc d) (𝒄ong' h t) = do + rT ← (L.map (λ (k , s) → R.strErr ("\n " <> mkNiceVar' "𝒙" k <> " = " <> trimLeft s )) + ∘S zipWithIndex) <$> (mapM (argRndr >=& ( indent' false ' ' 6)) t) + rHead ← inCuCtx ctx $ addNDimsToCtx' "𝒙" (length t) $ renderTerm h + pure $ "\nlet " ∷ₑ rT ++ "\nin " ∷ₑ [ rHead ]ₑ + + where + argRndr : Hco A B → R.TC _ + argRndr x = (((ppCT'' ctx d (hco' x)))) >>= R.formatErrorParts + + ppCT' : ℕ → CuTerm' A B → R.TC (List R.ErrorPart) + ppCT' = ppCT'' (defaultCtx dim) + + + +genAbstr : ℕ → String +genAbstr dim = "λ" <> + (L.foldl _<>_ "" $ L.map (λ k → (" " <> mkNiceVar' "𝓲" k)) (rev (range dim))) <> " → " + +codeGen : {A B : Type} (normaliseCells₁ : Bool) (dim : ℕ) → + ℕ → CuTerm' A B → R.TC String +codeGen nc dim fuel cu = ((genAbstr dim <>_) ∘S (indent' false ' ' 6)) <$> + (codeGen.ppCT' nc dim fuel cu >>= R.formatErrorParts) + +data ⁇ : Type where + +hcoFromIExpr : ℕ → FExpr → R.Term → R.TC CuTerm +hcoFromIExpr dim fe (R.def (quote ⁇) []) = + pure $ hco ((_, metaCell) <$> fe) metaCell +hcoFromIExpr dim fe tm' = do + let tm = liftVarsFrom dim zero tm' + xs ← mapM (λ sf → (sf ,_) <$> (cell ∘S liftVars <$> pure (subfaceCell sf tm)) ) fe + pure (hco xs (cell tm)) diff --git a/Cubical/Tactics/Reflection/Dimensions.agda b/Cubical/Tactics/Reflection/Dimensions.agda new file mode 100644 index 0000000000..adb0114a08 --- /dev/null +++ b/Cubical/Tactics/Reflection/Dimensions.agda @@ -0,0 +1,735 @@ +{- + +This module provides utilities for working with interval expressions (`IExpr`) in a reflected syntax. +It introduces an abstract representation of interval expressions and face expressions (`FExpr`). + +1. **Interval Expressions (IExpr)** + - Defined as `List (List (Bool × ℕ))`. + - Operations: + - `∨'` (disjunction) + - `∧'` (conjunction) + - `~'` (negation) + - `ie[ _ ]` (singleton interval expression) + - Normalization using `normIExpr`. + - Conversion between `IExpr` and terms (`IExpr→Term`). + +2. **Face Expressions (FExpr)** + - Defined as `List SubFace`. + - Operations: + - `fe∷` (face extension) + - `fe∷×` (face extension with additional data) + - `++fe` (concatenation for `FExpr`) + - `++fe×` (concatenation for `FExpr` with additional data) + - Conversion between `IExpr` and `FExpr` via `I→F`. + +3. **SubFaces and Constraints** + - Definition and manipulation of faces and subfaces in the interval (e.g., `allSubFacesOfDim`, `subFaceConstraints`). + - Operations and relations between subfaces (``, `sf∩`, ``). + +4. **Contextual Operations** + - Addition of interval dimensions to contexts (`addNDimsToCtx`). + - Handling and applying face constraints within contexts. + +5. **Utility Functions** + - Extraction and manipulation of interval expressions from terms (`extractIExpr`, `extractIExprM`). + - Lifting and dropping variables in terms based on boolean masks (`liftWhere`, `dropWhere`). + +6. **Non-Degenerate Faces** + - Check and filter for non-degenerate faces and expressions (`nonDegFExpr`, `isNonDegen`). + + +The current representations of interval and face expressions are not very type-safe. +While functions for manipulating and combining expressions typically produce normalized versions, +this is not enforced at the type level. Additionally, the representations are not parameterized +by context, meaning the well-scopedness of these expressions is not enforced. +Making these aspects type-safe should not be difficult, would be a natural progression and should be a focus +for future refactoring. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.Reflection.Dimensions where + +import Agda.Builtin.Reflection as R + +open import Agda.Builtin.String +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Reflection.Base +open import Cubical.Reflection.Sugar + +open import Cubical.Data.List as L +open import Cubical.Data.Bool as 𝟚 +open import Cubical.Data.Sigma +open import Cubical.Data.Sum +open import Cubical.Data.Nat +open import Cubical.Data.Empty +open import Cubical.Data.Maybe as Mb + +open import Cubical.Relation.Nullary + +open import Cubical.Tactics.Reflection.Variables +open import Cubical.Tactics.Reflection.Utilities +open import Cubical.Tactics.Reflection.Error + +IExpr : Type +IExpr = List (List (Bool × ℕ)) + +private + variable + ℓ : Level + A : Type ℓ + + + +_∨'_ : IExpr → IExpr → IExpr +_∨'_ = _++_ + + +_∧'_ : IExpr → IExpr → IExpr +x ∧' y = L.map (uncurry _++_) (cart x y) + + +~' : IExpr → IExpr +~' = foldl (λ x y → x ∧' L.map ([_] ∘S map-fst not) y) [ [] ] + +ie[_] : ℕ → IExpr +ie[ x ] = [ [ true , x ] ] + + + + +normIExpr : IExpr → IExpr +normIExpr = norm∨ ∘S L.map norm∧ + where + disc𝟚×ℕ = (discreteΣ 𝟚._≟_ λ _ → discreteℕ) + + norm∧ : List (Bool × ℕ) → List (Bool × ℕ) + norm∧ = nub disc𝟚×ℕ + + norm∨ : IExpr → IExpr + + norm∨' : List (Bool × ℕ) → IExpr → IExpr + norm∨' x [] = [ x ] + norm∨' x (x' ∷ xs) with subs? disc𝟚×ℕ x x' | subs? disc𝟚×ℕ x' x + ... | false | false = x' ∷ norm∨' x xs + ... | false | true = norm∨' x' xs + ... | true | b' = norm∨' x xs + + + + norm∨ [] = [] + norm∨ (x ∷ x₁) = norm∨' x x₁ + + +vlamⁿ : ℕ → R.Term → R.Term +vlamⁿ zero t = t +vlamⁿ (suc n) t = vlam "𝒊" (vlamⁿ n t) + + + +$i : ∀ {ℓ} {A : Type ℓ} → (I → A) → I → A +$i = λ f i → f i + +$≡ : ∀ {ℓ} {A : I → Type ℓ} {x : A i0} {y : A i1} → (PathP A x y) → ∀ i → A i +$≡ f i = f i + + +$PI : ∀ {ℓ} (A : Type ℓ) → (I → (Partial i1 A)) → I → A +$PI _ f i = f i 1=1 + + +appNDimsI : ℕ → R.Term → R.Term +appNDimsI zero t = t +appNDimsI (suc n) t = + appNDimsI n $ R.def (quote $i) ( t v∷ v[ R.var n [] ]) + +SubFace = List (Maybe Bool) + +allSubFacesOfDim : ℕ → List SubFace +allSubFacesOfDim zero = [ [] ] +allSubFacesOfDim (suc x) = + let l = allSubFacesOfDim x + in L.map (nothing ∷_) l + ++ L.map (just false ∷_) l + ++ L.map (just true ∷_) l + +sfDim : SubFace → ℕ +sfDim sf = length sf ∸ length (filter (λ { (just _) → true ; _ → false} ) sf) + +allSubFacesOfSfDim : ℕ → ℕ → List SubFace +allSubFacesOfSfDim n k = filter ((_=ℕ k) ∘S sfDim) $ allSubFacesOfDim n + +subFaceConstraints : SubFace → List (Bool × ℕ) +subFaceConstraints [] = [] +subFaceConstraints (x ∷ xs) = + Mb.rec (idfun _) (_∷_ ∘S (_, zero)) x $ + (L.map (map-snd suc) (subFaceConstraints xs)) + + + +injSF : SubFace → SubFace → SubFace +injSF [] y = [] +injSF (x ∷ x₁) [] = [] +injSF (nothing ∷ x₁) (x₂ ∷ y) = x₂ ∷ injSF x₁ y +injSF (just x ∷ x₁) (x₂ ∷ y) = injSF x₁ y + + +data : Type where + ⊂ ⊃ ⊃⊂ ⊂⊃ : + +⊂⊃? : → Bool +⊂⊃? ⊂⊃ = true +⊂⊃? _ = false + + + +SF⊃⊂ : → Type +SF⊃⊂ ⊃⊂ = Unit +SF⊃⊂ _ = ⊥ + +_<⊕>_ : +⊃⊂ <⊕> x₁ = ⊃⊂ +⊂⊃ <⊕> x₁ = x₁ + +⊂ <⊕> ⊂ = ⊂ +⊂ <⊕> ⊂⊃ = ⊂ +⊂ <⊕> _ = ⊃⊂ + +⊃ <⊕> ⊂⊃ = ⊃ +⊃ <⊕> ⊃ = ⊃ +⊃ <⊕> _ = ⊃⊂ + +__ : Maybe Bool → Maybe Bool → +nothing nothing = ⊂⊃ +nothing just x = ⊃ +just x nothing = ⊂ +just x just x₁ = if (x ⊕ x₁) then ⊃⊂ else ⊂⊃ + +__ : SubFace → SubFace → +[] [] = ⊂⊃ +[] (y ∷ ys) = (nothing y) <⊕> ([] ys) +(x ∷ xs) [] = (x nothing) <⊕> (xs []) +(x ∷ xs) (y ∷ ys) = (x y) <⊕> (xs ys) + +_sf∩_ : SubFace → SubFace → Maybe SubFace +[] sf∩ [] = just [] +[] sf∩ (x ∷ x₁) = nothing +(x ∷ x₁) sf∩ [] = nothing +(nothing ∷ x₁) sf∩ (x₂ ∷ x₃) = + map-Maybe (x₂ ∷_) (x₁ sf∩ x₃ ) +(just x ∷ x₁) sf∩ (nothing ∷ x₃) = map-Maybe (just x ∷_) (x₁ sf∩ x₃ ) +(just x ∷ x₁) sf∩ (just x₂ ∷ x₃) = + if (x ⊕ x₂) + then nothing + else map-Maybe (just x ∷_) (x₁ sf∩ x₃ ) + +FExpr = List SubFace + + + + + +infixr 5 _fe∷_ _fe∷×_ + +_fe∷_ : SubFace → FExpr → FExpr +x fe∷ [] = x ∷ [] +x fe∷ y@(sf ∷ x₁) with sf x +... | ⊂ = x fe∷ x₁ +... | ⊃ = y +... | ⊃⊂ = sf ∷ (x fe∷ x₁) +... | ⊂⊃ = y + + +_fe∷×_ : ∀ {ℓ} {A : Type ℓ} → (SubFace × A) → List (SubFace × A) → List (SubFace × A) +(x , a) fe∷× [] = (x , a) ∷ [] +(x , a) fe∷× y@((sf , a') ∷ x₁) with sf x +... | ⊂ = (x , a) fe∷× x₁ +... | ⊃ = y +... | ⊃⊂ = (sf , a') ∷ ((x , a) fe∷× x₁) +... | ⊂⊃ = y + + +_++fe_ : FExpr → FExpr → FExpr +x ++fe y = foldr _fe∷_ y x + +_++fe×_ : List (SubFace × A) → List (SubFace × A) → List (SubFace × A) +x ++fe× y = foldr _fe∷×_ y x + + +_⊂?_ : SubFace → FExpr → Bool +x ⊂? [] = false +x ⊂? (y ∷ sf) with x y +... | ⊂ = true +... | ⊃ = x ⊂? sf +... | ⊃⊂ = x ⊂? sf +... | ⊂⊃ = true + + +¬⊃⊂? : → Bool +¬⊃⊂? ⊃⊂ = false +¬⊃⊂? _ = true + + +mkFace : Bool → ℕ → SubFace +mkFace b k = iter k (nothing ∷_ ) (just b ∷ []) + +_∪Mb_ : Maybe Bool → Maybe Bool → Maybe (Maybe Bool) +nothing ∪Mb x₁ = just x₁ +just x ∪Mb nothing = just (just x) +just x ∪Mb just x₁ = if (x ⊕ x₁) then nothing else just (just x) + +_∪SF_ : SubFace → SubFace → Maybe SubFace +[] ∪SF x₁ = just x₁ +(x ∷ x₂) ∪SF [] = just ((x ∷ x₂)) +(x ∷ x₂) ∪SF (x₁ ∷ x₃) = + map2-Maybe _∷_ (x ∪Mb x₁) (x₂ ∪SF x₃) + +fromSF : (List (Bool × ℕ)) → Maybe SubFace +fromSF [] = nothing +fromSF (x ∷ xs) = + foldr (flip (Mb.rec (λ _ → nothing) _∪SF_ ) ∘S uncurry mkFace) (just (mkFace (fst x) (snd x))) xs + + +withAllIncluded : ℕ → FExpr → FExpr +withAllIncluded dim x = + filter (_⊂? x) (allSubFacesOfDim dim) + + +SubFace→Term : SubFace → R.Term +SubFace→Term = h 0 + where + h : ℕ → SubFace → R.Term + h k [] = R.con (quote i1) [] + h k (nothing ∷ xs) = h (suc k) xs + h k (just x ∷ xs) = + let x' = R.var k [] + x'' = if x then x' else R.def (quote ~_) v[ x' ] + in R.def (quote _∧_) (x'' v∷ v[ h (suc k) xs ]) + + +I→F : IExpr → FExpr +I→F [] = [] +I→F (x ∷ x₁) = + Mb.rec + (I→F x₁) + (λ x → x fe∷ I→F x₁) + (fromSF x) + + + +endTerm : Bool → R.Term +endTerm = (if_then R.con (quote i1) [] else R.con (quote i0) []) + + +IExpr∧→Term : List (Bool × ℕ) → R.Term +IExpr∧→Term [] = endTerm true +IExpr∧→Term ((b , k) ∷ []) = (let x' = R.var (k) [] in if b then x' else R.def (quote ~_) v[ x' ]) +IExpr∧→Term ((b , k) ∷ xs) = + R.def (quote _∧_) + ((let x' = R.var (k) [] in if b then x' else R.def (quote ~_) v[ x' ]) v∷ v[ IExpr∧→Term xs ]) + + +getMaxVar : IExpr → ℕ +getMaxVar = maximum ∘S L.map snd ∘S join + where + maximum : List ℕ → ℕ + maximum = foldr max zero + + + +IExpr→Term : IExpr → R.Term +IExpr→Term [] = endTerm false +IExpr→Term (xs ∷ []) = IExpr∧→Term xs +IExpr→Term (xs ∷ xxs) = R.def (quote _∨_) (IExpr∧→Term xs v∷ v[ IExpr→Term xxs ]) + + +mapVarsInIExpr : (ℕ → ℕ) → IExpr → IExpr +mapVarsInIExpr f = L.map (L.map (map-snd f)) + +interpolateIExpr : IExpr → IExpr → IExpr → IExpr +interpolateIExpr t ie0 ie1 = normIExpr $ + (((~' t) ∧' ie0) ∨' ((t) ∧' ie1)) ∨' (ie0 ∧' ie1) + +IExpr→MaybeEnd : IExpr → Maybe Bool +IExpr→MaybeEnd = (λ { [] -> just false ; ([] ∷ []) → just true ; _ → nothing } ) ∘S normIExpr + +evalIExprOnFace : SubFace → IExpr → IExpr +evalIExprOnFace sf = catMaybes ∘S L.map (h (zipWithIndex sf)) + + where + + hh : ℕ × Maybe Bool → Maybe (List (Bool × ℕ)) → Maybe (List (Bool × ℕ)) + hh (_ , nothing) l = l + hh _ nothing = nothing + hh (k , just b) (just xs) = + if elem? (discreteΣ 𝟚._≟_ λ _ → discreteℕ) (not b , k) xs + then nothing + else just (filter (not ∘S Dec→Bool ∘ (discreteΣ 𝟚._≟_ λ _ → discreteℕ) (b , k)) xs) + + h : List (ℕ × Maybe Bool) → (List (Bool × ℕ)) → Maybe (List (Bool × ℕ)) + h sf xs = foldr hh (just xs) sf + + + + +module _ (s : String) where + + addNDimsToCtx' : ℕ → R.TC A → R.TC A + addNDimsToCtx' zero f = f + addNDimsToCtx' (suc k) = + R.extendContext (mkNiceVar' s k) (varg (R.def (quote I) [])) + ∘S (addNDimsToCtx' k) + +addNDimsToCtx = addNDimsToCtx' "𝓲" + + + +-- works under assumption, that all free variables in term are of Interval type +-- (wont work for term like (f 1) where (f : ℕ → I)) + +extractIExpr : R.Term → Maybe IExpr +extractIExpr (R.var x []) = ⦇ ( ie[ x ] ) ⦈ +extractIExpr (R.con (quote i0) []) = ⦇ [] ⦈ +extractIExpr (R.con (quote i1) []) = ⦇ ([ [] ]) ⦈ +extractIExpr (R.def (quote _∨_) (x v∷ v[ y ])) = + ⦇ extractIExpr x ∨' extractIExpr y ⦈ +extractIExpr (R.def (quote _∧_) (x v∷ v[ y ])) = + ⦇ extractIExpr x ∧' extractIExpr y ⦈ +extractIExpr (R.def (quote ~_) v[ x ] ) = + ⦇ ~' (extractIExpr x) ⦈ +extractIExpr _ = nothing + + +extractIExprM : R.Term → R.TC IExpr +extractIExprM t = Mb.rec (R.typeError ("extractIExpr fail:" ∷ₑ [ t ]ₑ )) pure (extractIExpr t) + +CuCtx = List (String × Maybe Bool) + +defaultCtx : ℕ → CuCtx +defaultCtx dim = L.map ((_, nothing) ∘S mkNiceVar' "𝓲") $ range dim + +inCuCtx : CuCtx → R.TC A → R.TC A +inCuCtx ctx x = foldl (λ x → uncurry + λ v → Mb.caseMaybe (R.extendContext v ((varg (R.def (quote I) []))) x) x) x ctx + +inCuCtx' : CuCtx → R.TC A → R.TC A +inCuCtx' ctx x = foldl (λ x → uncurry + λ v _ → R.extendContext v ((varg (R.def (quote I) []))) x) x ctx + + + + + +SubFace→TermInCtx : CuCtx → SubFace → R.Term +SubFace→TermInCtx ctx = h 0 + where + + kInCtx : CuCtx → ℕ → ℕ + kInCtx [] k = k + kInCtx ((fst₁ , nothing) ∷ xs) zero = zero + kInCtx ((fst₁ , nothing) ∷ xs) (suc k) = suc (kInCtx xs k) + kInCtx ((fst₁ , just x) ∷ xs) k = suc (kInCtx (xs) k) + + h : ℕ → SubFace → R.Term + h k [] = R.con (quote i1) [] + h k (nothing ∷ xs) = h (suc k) xs + h k (just x ∷ xs) = + let x' = R.var (kInCtx ctx k) [] + x'' = if x then x' else R.def (quote ~_) v[ x' ] + in R.def (quote _∧_) (x'' v∷ v[ h (suc k) xs ]) + +IExpr→TermInCtx : CuCtx → IExpr → R.Term +IExpr→TermInCtx ctx = + foldl (λ x y → R.def (quote _∨_) (x v∷ + v[ foldl (λ x (b , k) → R.def (quote _∧_) (x v∷ + v[ (let x' = R.var (kInCtx ctx k) [] + x'' = if b then x' else R.def (quote ~_) v[ x' ] + in x'') + ] )) + (R.con (quote i1) []) y ] )) + (R.con (quote i0) []) + where + + kInCtx : CuCtx → ℕ → ℕ + kInCtx [] k = k + kInCtx ((fst₁ , nothing) ∷ xs) zero = zero + kInCtx ((fst₁ , nothing) ∷ xs) (suc k) = suc (kInCtx xs k) + kInCtx ((fst₁ , just x) ∷ xs) k = suc (kInCtx (xs) k) + + +applyFaceConstraints : SubFace → CuCtx → CuCtx +applyFaceConstraints sf [] = [] +applyFaceConstraints [] ctx@(_ ∷ _) = ctx +applyFaceConstraints sf@(_ ∷ _) (c@(_ , just _) ∷ ctx) = + c ∷ applyFaceConstraints sf ctx +applyFaceConstraints (mbC ∷ sf) ((v , nothing) ∷ ctx) = + (v , mbC) ∷ applyFaceConstraints sf ctx + + +freeVars : CuCtx → List String +freeVars = L.map fst ∘S filter (λ { (_ , (nothing)) → true ; _ → false} ) + +boundedDims : CuCtx → List (String × Bool) +boundedDims = catMaybes ∘S L.map (uncurry λ s → map-Maybe (s ,_)) + + + +dropWhere : List Bool → R.Term → R.Term +dropWhere = dw ∘S rev + where + dw : List Bool → R.Term → R.Term + dw [] t = t + dw (b ∷ xs) t = dw xs (if b then dropVar (length xs) t else t) + +liftWhere : List Bool → R.Term → R.Term +liftWhere = lw ∘S rev + where + lw : List Bool → R.Term → R.Term + lw [] t = t + lw (b ∷ xs) t = + let t' = lw xs t + in (if b then (liftVarsFrom 1 (length xs) t') else t') + + +subfaceCell : SubFace → R.Term → R.Term +subfaceCell sf = dropWhere (L.map (λ {(just _) → true ; _ → false }) sf) ∘S replaceVarWithCon (r sf) + where + r : SubFace → ℕ → Maybe R.Name + r sf = map-Maybe (if_then quote i1 else quote i0) ∘S lookupAlways nothing sf + +subfaceCellNoDrop : SubFace → R.Term → R.Term +subfaceCellNoDrop sf = replaceVarWithCon (r sf) + where + r : SubFace → ℕ → Maybe R.Name + r sf = map-Maybe (if_then quote i1 else quote i0) ∘S lookupAlways nothing sf + + + + +feMask : FExpr → List Bool +feMask = foldr (alwaysZipWith + (λ x y → (fromJust-def false x) or (fromJust-def false y)) ∘S L.map (caseMaybe false true)) [] + + + + +degenRoughPrecheck : ℕ → IExpr → Bool +degenRoughPrecheck n xs = + let h = flip (elem? (discreteΣ 𝟚._≟_ λ _ → discreteℕ)) (join xs) in + foldr + (_or_ ∘S λ k → (h (true , k)) and (h (false , k))) + false (range n) + +F→I : ℕ → FExpr → IExpr +F→I n = + L.map (catMaybes ∘S L.map (λ (k , mbB) → map-Maybe (_, k) mbB) ∘S zipWithIndex) + + +allVertsOfSF : SubFace → List SubFace +allVertsOfSF sf = filter + (λ sf' → Dec→Bool (discreteℕ (sfDim sf') 0) and (sf' ⊂? [ sf ])) + (allSubFacesOfDim (length sf)) + + +undegen' : ℕ → IExpr → List (SubFace × List (List (Bool × ℕ)) × String × Bool) +undegen' n ie' = + + let ie = normIExpr ie' + h sf = sf , normIExpr (evalIExprOnFace sf ie) , + foldr _<>_ "" (L.map + (λ l → "[" <> foldr _<>_ "" (L.map (λ (b , k) → + let vv = mkNiceVar' "𝓲" k + in "(" <> (if b then vv else "~" <> vv) <> ")" ) l) <> "]" + + ) + (normIExpr (evalIExprOnFace sf ie))) , + let sfs = allVertsOfSF sf + in foldr + _and_ + true (L.map (λ sf' → Mb.fromJust-def false (IExpr→MaybeEnd (evalIExprOnFace sf' ie))) sfs) + + + in L.map h (allSubFacesOfDim n) + + +instance + ToErrorPartSubFace : ToErrorPart SubFace + ToErrorPart.toErrorPart ToErrorPartSubFace sf = + R.strErr (foldr (λ (i , mBb) s → + "(" <> Mb.rec "__" (if_then "i1" else "i0") mBb <> ")" <> s) + "" (zipWithIndex sf) ) + + + + +unifyTest : ℕ → R.Term → R.Term → R.TC Bool +unifyTest dim t t' = addNDimsToCtx dim $ + (R.unify t t' >> pure true) <|> (pure false) + + + + + + +nonDegFExpr : ℕ → IExpr → FExpr +nonDegFExpr n ie' = + let ie = normIExpr ie' + h sf = + let sfs = allVertsOfSF sf + in foldr + _and_ + true (L.map (λ sf' → Mb.fromJust-def false (IExpr→MaybeEnd (evalIExprOnFace sf' ie))) sfs) + in foldr _fe∷_ [] (filter h (allSubFacesOfDim n)) + + +undegen : Bool → ℕ → IExpr → IExpr +undegen onEnd n ie' = + + let ie = normIExpr ie' + w = nonDegFExpr n ie' + + in if onEnd + then (F→I n w) + else interpolateIExpr ie[ n ] ie (F→I n w) + +undegen2 : Bool → ℕ → IExpr → IExpr +undegen2 onEnd n ie' = + + let ie = normIExpr ie' + w = nonDegFExpr n ie' + + in if onEnd + then (F→I n w) + else interpolateIExpr ie[ zero ] (mapVarsInIExpr suc (F→I n w)) (mapVarsInIExpr suc ie) + + + +isNonDegen : ℕ → IExpr → R.TC Bool +isNonDegen dim iexpr = + unifyTest dim (IExpr→Term iexpr) (IExpr→Term (undegen true dim iexpr)) + + +undegenFcs : ℕ → List IExpr → (R.TC (FExpr)) +undegenFcs dim l = do + do + foldrM (λ sf fe → _++fe fe <$> (if ((sfDim sf) =ℕ 0) then pure [ sf ] else do + isNonDegForEvery ← foldrM + (\ie b → (b and_) <$> (isNonDegen dim (evalIExprOnFace sf ie))) + true l + pure $ if isNonDegForEvery then [ sf ] else [])) + [] (filter ((_<ℕ dim) ∘ sfDim) (allSubFacesOfDim dim)) + + +normIExprInTerm : ℕ → R.Term → R.TC R.Term +normIExprInTerm offset = + atVarOrDefM.rv + (λ n k _ args → R.var (n + k) <$> args) + h + zero + + where + + g : R.Name → List (R.Arg R.Term) → R.Term → Maybe R.Term + g (quote _∨_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote _∧_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote ~_) a@(v[ _ ]) tm = just tm + g _ _ _ = nothing + + h : ℕ → + R.Name → + List (R.Arg R.Term) → R.TC (List (R.Arg R.Term)) → R.TC R.Term + h _ nm arg argM = + Mb.rec (R.def nm <$> argM) + ((extractIExprM >=& + (IExpr→Term + ∘ mapVarsInIExpr (_+ offset) + ∘ normIExpr + ∘ mapVarsInIExpr (_∸ offset) ))) + (g nm arg (R.def nm arg)) + + +macro + normIExprInTermM : R.Term → R.Term → R.TC Unit + normIExprInTermM t h = + normIExprInTerm zero t >>= flip R.unify h + + + +extractAllIExprs : R.Term → List IExpr +extractAllIExprs tm = + snd $ runIdentity $ unwrap (atVarOrDefM.rv {M = [ State₀T (List IExpr) RMT IdentityF ]_ } + (λ _ v _ argM → R.var v <$> argM) + gg zero tm) [] + where + + g : R.Name → List (R.Arg R.Term) → Bool + g (quote _∨_) a@(_ v∷ v[ _ ]) = true + g (quote _∧_) a@(_ v∷ v[ _ ]) = true + g (quote ~_) a@(v[ _ ]) = true + g _ _ = false + + + gg : _ + gg n nm arg argM = let t = R.def nm arg in + if (g nm arg) + then (Mb.rec (liftM (identity tt)) + (λ ie → modify ((mapVarsInIExpr (_∸ n) ie) ∷_)) (extractIExpr t) ) >> pure t + else R.def nm <$> argM + +mapIExprs : ℕ -> ℕ → (IExpr → IExpr) → R.Term → R.TC R.Term +mapIExprs dim offset fn = + atVarOrDefM.rv + (λ n k _ args → + if (n <ℕ (suc k)) and (k <ℕ (n + dim)) + then (pure ((IExpr→Term + ∘S mapVarsInIExpr (_+ (offset + n)) + ∘S fn + ∘S mapVarsInIExpr (_∸ (offset + n))) (ie[ k ]))) + else (R.var k <$> args)) + h + zero + + where + + g : R.Name → List (R.Arg R.Term) → R.Term → Maybe R.Term + g (quote _∨_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote _∧_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote ~_) a@(v[ _ ]) tm = just tm + g _ _ _ = nothing + + h : ℕ → + R.Name → + List (R.Arg R.Term) → R.TC (List (R.Arg R.Term)) → R.TC R.Term + h n nm arg argM = + Mb.rec (R.def nm <$> argM) + ((extractIExprM >=& + (IExpr→Term + ∘ mapVarsInIExpr (_+ (offset + n)) + ∘ fn + ∘ mapVarsInIExpr (_∸ (offset + n)) ))) + (g nm arg (R.def nm arg)) + + +icConnFree' : IExpr → Bool +icConnFree' [] = true +icConnFree' ([] ∷ []) = true +icConnFree' ((x ∷ []) ∷ []) = true +icConnFree' _ = false + +icConnFree : IExpr → Bool +icConnFree = icConnFree' ∘ normIExpr + +missingSubFaces : ℕ → FExpr → List SubFace +missingSubFaces dim fe = filter (not ∘S _⊂? fe) (allSubFacesOfDim dim) + +getCuCtx : R.TC CuCtx +getCuCtx = + takeWhile + (λ { (s , varg (R.def (quote I) [])) → just (s , nothing) + ; _ → nothing }) + <$> R.getContext diff --git a/Cubical/Tactics/Reflection/Error.agda b/Cubical/Tactics/Reflection/Error.agda new file mode 100644 index 0000000000..72bcbd6c29 --- /dev/null +++ b/Cubical/Tactics/Reflection/Error.agda @@ -0,0 +1,222 @@ +{- +This module provides utilities designed for printing messages and errors in the context of macros. + +Key functionalities include: + +- **Error Message Composition:** + - Custom composable error parts using the `_∷ₑ_`, `_++ₑ_`, `_∷nl_`, and `_++nl_` operators. + - Instances for converting various types (`String`, `Char`, `ℕ`, `Bool`, `R.Term`, `R.Name`, `R.ErrorPart`) to `R.ErrorPart`. + +- **String Manipulation:** + - Indentation and offsetting functions (`indent'`, `indent`, `indentₑ`, `offsetStr`, `offsetStrR`). + - String formatting and line handling (`lines`). + +- **Error Rendering:** + - Functions for rendering terms and arguments (`renderTerm`, `renderArg`). + - Concatenating runs of consecutive strErr in List of ErrorParts (`<>StrErr`). + +- **Result Wrappers:** + - Wrapping results and errors in the `ResultIs` type using `wrapResult` and `wrapError`. + +- **Testing Helpers:** + - `assertNoErr` function to facilitate writing tests that check for the presence or absence of errors. + - `TestResult` data type which includes `✓-pass` and `⊘-fail` constructors. + +- **Macros:** + - `showCtx` macro for printing the current context in the form of a type error. + - `showTeles` function to generate a list of error parts from a telescope. + +-} + +{-# OPTIONS --safe #-} + +module Cubical.Tactics.Reflection.Error where + + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String +open import Agda.Builtin.Char + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Reflection.Base +open import Cubical.Reflection.Sugar + +open import Cubical.Data.List +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Sigma + +open import Cubical.Tactics.Reflection.Variables +open import Cubical.Tactics.Reflection.Utilities + +record ToErrorPart {ℓ} (A : Type ℓ) : Type ℓ where + field + toErrorPart : A → R.ErrorPart + +open ToErrorPart + +infixr 5 _∷ₑ_ _∷nl_ _++ₑ_ _++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) ++_ + +[_]ₑ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart +[_]ₑ = _∷ₑ [] + +instance + ToErrorPartString : ToErrorPart String + toErrorPart ToErrorPartString = R.strErr + + ToErrorPartChar : ToErrorPart Char + toErrorPart ToErrorPartChar = R.strErr ∘S primStringFromList ∘S [_] + + + ToErrorPartℕ : ToErrorPart ℕ + toErrorPart ToErrorPartℕ = R.strErr ∘ primShowNat + + ToErrorPartBool : ToErrorPart Bool + toErrorPart ToErrorPartBool = R.strErr ∘ (if_then "𝟙" else "𝟘") + + + ToErrorPartTerm : ToErrorPart R.Term + toErrorPart ToErrorPartTerm = R.termErr + + ToErrorPartName : ToErrorPart R.Name + toErrorPart ToErrorPartName = R.nameErr + + ToErrorPartErrorPart : ToErrorPart R.ErrorPart + toErrorPart ToErrorPartErrorPart x = x + + +_∷nl_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → A → List R.ErrorPart → List R.ErrorPart +_∷nl_ x y = x ∷ₑ "\n" ∷ₑ y + +_++nl_ : ∀ {ℓ} {A : Type ℓ} → {{ToErrorPart A}} → List A → List R.ErrorPart → List R.ErrorPart +_++nl_ x y = x ++ₑ "\n" ∷ₑ y + + +<>StrErr : List R.ErrorPart → List R.ErrorPart +<>StrErr [] = [] +<>StrErr (x ∷ xs) = h x xs + where + h : R.ErrorPart → List R.ErrorPart → List R.ErrorPart + h x [] = [ x ] + h (R.strErr x) (R.strErr y ∷ xs) = h (R.strErr (x <> y)) xs + h x (y ∷ xs) = x ∷ h y xs + +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 + + + +unArgs : List (R.Arg (R.Term)) → List R.ErrorPart +unArgs [] = [] +unArgs (R.arg i x ∷ x₁) = x ∷ₑ unArgs x₁ + +renderTerm : R.Term → R.TC String +renderTerm = R.formatErrorParts ∘ [_]ₑ + +renderArg : R.Arg R.Term → R.TC String +renderArg (R.arg i x) = renderTerm x + + +stringLength : String → ℕ +stringLength = length ∘S primStringToList + + +indent' : Bool → Char → ℕ → String → String +indent' _ _ zero x = x +indent' b ch k = + primStringFromList ∘S + (if b then ((ch ∷_) ∘S (ind ++_)) else (idfun _)) ∘S h ∘S primStringToList + + where + ind = repeat k ' ' + + h : List Char → List Char + h [] = [] + h (x ∷ xs) with primCharEquality x '\n' + ... | false = x ∷ h xs + ... | true = '\n' ∷ ch ∷ ind ++ h xs + + +indent = indent' true + +indentₑ : ℕ → List R.ErrorPart → List R.ErrorPart +indentₑ k = map (λ { (R.strErr x) → R.strErr (indent ' ' k x) ; x → x }) ∘S <>StrErr + +offsetStr : ℕ → String → String +offsetStr k = primStringFromList ∘S offset ' ' k ∘S primStringToList + +offsetStrR : ℕ → String → String +offsetStrR k = primStringFromList ∘S offsetR ' ' k ∘S primStringToList + +trimLeft : String → String +trimLeft = + primStringFromList + ∘S dropBy (λ { ' ' → true ; '\n' → true ; '\t' → true ; _ → false } ) + ∘S primStringToList + +data ResultIs {ℓ} {A : Type ℓ} : A → Type ℓ where + resultIs : ∀ s → ResultIs s + + +wrapResult : ∀ {ℓ} {A : Type ℓ} → R.Term → A → R.TC Unit +wrapResult hole x = do + x' ← R.quoteTC x + R.unify (R.con (quote resultIs) v[ x' ]) hole + + +lines : String → List String +lines = map primStringFromList ∘S h [] ∘S primStringToList + where + h : List (List Char) → List Char → List (List Char) + h xxs [] = xxs + h xxs ('\n' ∷ xs) = xxs ++ h [] xs + h [] (x ∷ xs) = h [ [ x ] ] xs + + h xxs (x ∷ xs) = h ((init xxs) ++ map (_∷ʳ x) (drop (predℕ (length xxs)) xxs)) xs + + + +wrapError : R.Term → List (R.ErrorPart) → R.TC Unit +wrapError hole x = do + x' ← ((map (offsetStrR 45) ∘S lines) <$> R.formatErrorParts x) >>= R.quoteTC + R.unify (R.con (quote resultIs) v[ x' ]) hole + + +data TestResult : Type where + ✓-pass ⊘-fail : TestResult + +assertNoErr : ∀ {ℓ} {A : Type ℓ} → R.Term → R.TC A → R.TC Unit +assertNoErr h x = do + (x >> wrapResult h ✓-pass) <|> wrapResult h ⊘-fail + + +visibillityWrap : R.Visibility → String → String +visibillityWrap R.visible x = " " <> x <> " " +visibillityWrap R.hidden x = "{" <> x <> "}" +visibillityWrap R.instance′ x = "⦃" <> x <> "⦄" + +showTeles : R.Telescope → R.TC (List R.ErrorPart) +showTeles = concatMapM h ∘S liftedTele + where + h : String × R.Arg R.Type → R.TC (List R.ErrorPart) + h (s , R.arg (R.arg-info v m) ty) = do + pure $ s ∷ₑ " : " ∷ₑ ty ∷nl [] + +macro + showCtx : R.Term → R.TC Unit + showCtx _ = R.getContext >>= (showTeles >=> R.typeError) + diff --git a/Cubical/Tactics/Reflection/QuoteCubical.agda b/Cubical/Tactics/Reflection/QuoteCubical.agda new file mode 100644 index 0000000000..86d367d419 --- /dev/null +++ b/Cubical/Tactics/Reflection/QuoteCubical.agda @@ -0,0 +1,381 @@ +{- + +## Cubical.Tactics.Reflection.QuoteCubical + +This module provides functionality for reflecting cubical terms within Agda, focusing on handling higher-dimensional compositions. It leverages the capabilities defined in `Cubical.Tactics.Reflection.CuTerm` to extract cubical terms. + +### Main Types and Functions: + +* `MarkHCompsVar`: + * A marker datatype used internally to handle `hcomp` terms during reflection. + +* `traceHComps`: + * Given a term, it traces and encodes `hcomp` terms, substituting them with a marker. + * Returns either a string message in case of error or a pair of the transformed term and the list of identified `hcomp` terms. + +#### Internal Monad `M`: + * Composed of states and error management, used to encapsulate reflection operations. + +#### Functions for Encoding and Decoding `hcomp` Terms: + * `codeHcomps`: Encodes `hcomp` terms by replacing them with a marker (`markHCompsVar`). + * `decodeHcomps`: Decodes `hcomp` markers back to their original term representations. + + +### Extraction and Evaluation: + +* `extractCuTerm'`: + * Extracts cubical terms recursively, handling `hcomp` and `𝒄ong` constructions based on specified depth and dimension parameters. + * Uses auxiliary functions like `checkHcomp` to handle `hcomp` terms and `try𝒄ong` to manage `𝒄ong` terms. + +* `quoteCuTerm` and `extractCuTerm`: + * Wrappers around `extractCuTerm'` to initiate extraction with a default depth of 100. + +* `pathApp` and `pathAppN`: + * Utility functions to apply a term along a path, handling various term constructions recursively. + +### Boundary Matching and Normalization: + +* `matchNCubeF`, `matchNCubeP`, and `matchNCube`: + * Match types od higher-dimensional cubes (n-cubes) and reduce them to their boundary terms. + +* `NBoundaryTerm`: + * Represents terms at the boundaries of cubes, encapsulating a term and its list of boundaries. + +* `quoteBd` and `quoteBdNC`: + * Quote the boundaries of n-cubes to CuTerm representation. + + +### Utilities for Extracting and Displaying Cubical Terms: + +* `extractCuTermFromNPath`: Extracts cubical terms from n-path terms, calculating the appropriate dimension. + +* `showCuTerm` and `showCuCode` (macros): + * Display structured or code representation of cubical terms for a given term and its inferred type. + + +-} + +{-# OPTIONS --safe #-} +module Cubical.Tactics.Reflection.QuoteCubical where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + + +open import Cubical.Data.Bool +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.List as L +open import Cubical.Data.Nat as ℕ +open import Cubical.Data.Sigma +open import Cubical.Data.Nat.Order.Recursive as ℕOR +open import Cubical.Data.Empty + +open import Cubical.Foundations.Univalence using (ua) + +open import Cubical.Reflection.Base renaming (v to 𝒗) +open import Cubical.Reflection.Sugar +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.Reflection.Utilities + +open import Cubical.Tactics.PathSolver.Reflection +open import Cubical.Tactics.Reflection.Error + +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.CuTerm + +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) + +data MarkHCompsVar : Type where + markHCompsVar : MarkHCompsVar + +traceHComps : R.Term → String ⊎.⊎ (R.Term × List R.Term) +traceHComps t = + let (x , xs) = runIdentity (unwrap (unwrap (codeHcomps t)) []) + in ⊎.map (idfun _) + ((_, xs) ∘S decodeHcomps (length xs) ∘S liftVarsFrom (length xs) 0) x + + where + + M : Functorω + M = [ Plus₀T String RMT [ State₀T (List R.Term) RMT IdentityF ]_ ]_ + + + codeHcomps : R.Term → M (R.Term) + codeHcomps = atVarOrDefM.rv + (λ n k args argsM → R.var k <$> argsM) + atDef + zero + where + atDef : ℕ → R.Name → (List (R.Arg R.Term)) → M (List (R.Arg R.Term)) → M R.Term + atDef n (quote hcomp) args _ = do + s ← liftM get + liftM (modify (_∷ʳ R.def (quote hcomp) args)) + pure (R.def (quote markHCompsVar) v[ R.lit (R.nat (length {A = R.Term} s)) ]) + atDef _ nm _ argsM = R.def nm <$> argsM + + decodeHcomps : ℕ → R.Term → R.Term + decodeHcomps m = runIdentity ∘S + atVarOrDefM.rv + (λ _ v _ → R.var v <$>_) + (λ n → λ where + (quote markHCompsVar) v[ R.lit (R.nat v) ] _ → pure $ R.var (v + n) [] + nm _ → R.def nm <$>_) zero + +normaliseWithType : R.Type → R.Term → R.TC R.Term +normaliseWithType ty tm = R.withNormalisation true $ R.checkType tm ty + +module ECT where + + evalOp : (Maybe R.Type) → ℕ → R.Term → R.TC R.Term + evalOp mbTy dims = + Mb.rec R.normalise (normaliseWithType ∘S liftVarsFrom dims zero) mbTy + + + getCuCase : R.Term → R.TC (Maybe ((R.Term × R.Term × R.Term × R.Term) × IExpr)) + getCuCase (R.def (quote hcomp) ( _ h∷ A h∷ φTm h∷ fcs v∷ v[ cap ])) = do + R.debugPrint "getCuCaseφ" 5 $ "getCuCase' φ :" ∷ₑ [ φTm ]ₑ + (just ∘ ((A , φTm , fcs , cap) ,_)) <$> extractIExprM φTm + getCuCase _ = pure nothing + + + try𝒄ong : ℕ → ℕ → R.Term → R.TC (Maybe (R.Term × List (CuTerm))) + + checkHcomp : Maybe R.Type → ℕ → ℕ → R.Term → R.Term → R.Term → R.Term → FExpr → R.TC CuTerm + + extractCuTerm' : Maybe R.Type → ℕ → ℕ → R.Term → R.TC CuTerm + + checkHcomp mbTy zero _ _ _ _ _ _ = R.typeError [ "checkHcomp FAIL : max depth" ]ₑ + checkHcomp mbTy (suc m) dim A φTm fcs lid φ = do + ⦇ hco + (mapM (λ sf → do + let tmA = subfaceCell sf fcs + Atm = subfaceCell sf A + -- tmA' = replaceVarWithCon (λ { zero → just (quote i0) ; _ → nothing }) fcs + tmB = (R.def (quote $PI) ((liftVars Atm) v∷ ((liftVars tmA)) + v∷ v[ R.var zero [] ] )) + sfbo = tmB + + ⦇ ⦇ sf ⦈ , (extractCuTerm' mbTy m (suc (sfDim sf)) sfbo) ⦈ + ) φ) + ((addNDimsToCtx dim (evalOp mbTy dim lid)) >>= extractCuTerm' mbTy m dim) ⦈ + + + + + try𝒄ong zero _ _ = R.typeError [ "extractCuTerm FAIL : max depth" ]ₑ + + try𝒄ong (suc m) dim t = do + ⊎.rec (R.typeError ∘S [_]ₑ) + withTracing + (traceHComps t) + where + + withTracing : R.Term × List R.Term → R.TC (Maybe (R.Term × List (CuTerm))) + withTracing (_ , []) = pure nothing + withTracing (R.var zero [] , _) = pure nothing + withTracing (h , tl) = ⦇ just ⦇ ⦇ h ⦈ , (mapM (extractCuTerm' nothing m dim) tl) ⦈ ⦈ + + extractCuTerm' mbTy zero _ _ = R.typeError [ "extractCuTerm FAIL : max depth" ]ₑ + extractCuTerm' mbTy (suc m) dim t = do + t ← addNDimsToCtx dim $ evalOp mbTy dim t + + addNDimsToCtx dim (getCuCase t) >>= + Mb.rec ( (pure t ) + >>= λ t' → --R.debugPrint "checkHcomp" 4 ("cell: \n " ∷ₑ [ tt ]ₑ) >> + ( try𝒄ong m dim t' ) >>= pure ∘S Mb.rec (cell t') (uncurry 𝒄ongF) + ) + λ ((A , φTm , fcs , cap) , φ) → + (checkHcomp + mbTy + -- (just (Mb.fromMaybe A mbTy)) + m dim A φTm fcs cap (L.map (offsetR nothing dim) (I→F φ))) + + + +quoteCuTerm : Maybe R.Type → ℕ → R.Term → R.TC CuTerm +quoteCuTerm = flip ECT.extractCuTerm' 100 + +extractCuTerm : Maybe R.Type → ℕ → R.Term → R.TC CuTerm +extractCuTerm mbTy dim = ECT.extractCuTerm' mbTy 100 dim ∘S appNDimsI dim ∘S liftVarsFrom dim 0 + + + +pathApp : R.Term → R.Term +pathApp (R.var x₁ args) = R.var (suc x₁) $ (LiftFrom.ra 1 0 args) ++ v[ 𝒗 zero ] +pathApp (R.con c args) = R.con c $ (LiftFrom.ra 1 0 args) ++ v[ 𝒗 zero ] +pathApp (R.def f args) = R.def f $ (LiftFrom.ra 1 0 args) ++ v[ 𝒗 zero ] +pathApp (R.lam v (R.abs _ t)) = t +pathApp (R.pat-lam cs args) = + R.pat-lam (LiftFrom.rc 1 0 cs) $ (LiftFrom.ra 1 0 args) ++ v[ 𝒗 zero ] +pathApp (R.pi a b) = R.lit (R.string ("unexpected in pathApp")) +pathApp (R.agda-sort s) = R.lit (R.string ("unexpected in pathApp")) +pathApp (R.lit l) = R.lit (R.string ("unexpected in pathApp")) +pathApp (R.meta x₁ args) = R.meta x₁ $ (LiftFrom.ra 1 0 args) ++ v[ 𝒗 zero ] +pathApp R.unknown = R.unknown + +extractCuTermFromPath : Maybe R.Type → R.Term → R.TC CuTerm +extractCuTermFromPath mbTy = ECT.extractCuTerm' mbTy 100 1 ∘S pathApp + + +pathAppN : ℕ → R.Term → R.Term +pathAppN = flip iter pathApp + +NBoundaryTerm : Type +NBoundaryTerm = (R.Term × List (R.Term × R.Term)) + + +matchNCubeF : ℕ → R.Term → R.TC NBoundaryTerm +matchNCubeF (suc fuel) (R.def (quote PathP) (l h∷ T v∷ x v∷ y v∷ [])) = do + (T' , xs) ← + (addNDimsToCtx 1 $ + (do T' ← R.normalise $ appNDimsI 1 (liftVars T) + matchNCubeF fuel T')) + let d = (length xs) + T0 = replaceAtTrm d (endTerm false) T' + T1 = replaceAtTrm d (endTerm true) T' + -- let + x' ← (addNDimsToCtx d $ normaliseWithType (T0) $ pathAppN d x) -- <|> R.typeError [ "matchNCubeF T0" ]ₑ + y' ← (addNDimsToCtx d $ normaliseWithType (T1) $ pathAppN d y) -- <|> R.typeError [ "matchNCubeF T1" ]ₑ + pure (T' , (x' , y') ∷ xs) +matchNCubeF _ tm = (pure $ tm , []) + +matchNCubeP matchNCube : R.Term → R.TC NBoundaryTerm +matchNCubeP = matchNCubeF 20 + +matchNCube x = do + (A , fcs) ← matchNCubeF 20 x + if (hasVarBy (_<ℕ length fcs) A) then R.typeError [ "not a _≡_ " ]ₑ else + pure ((dropVars.rv (length fcs) zero A , fcs)) + +quoteBd : NBoundaryTerm → R.TC CuBoundary +quoteBd (A , xs) = do + let dim = predℕ (length xs) + mapM (λ (t0 , t1) → ⦇ quoteCuTerm (just A) dim t0 , quoteCuTerm (just A) dim t1 ⦈ ) xs + + +matchSquare : NBoundaryTerm → Maybe (R.Term × ((R.Term × R.Term)×(R.Term × R.Term))) +matchSquare (A , ((a₀₋ , a₁₋) ∷ (a₋₀ , a₋₁) ∷ [])) = + just (A , ((a₀₋ , a₁₋) , (a₋₀ , a₋₁))) +matchSquare _ = nothing + +unquoteNCube : NBoundaryTerm → R.Type +unquoteNCube (A , []) = A +unquoteNCube (A , (f0 , f1) ∷ fcs) = + let t = unquoteNCube (A , fcs) + in R.def (quote PathP) + (vlam "𝒊" t v∷ vlamⁿ (length fcs) f0 v∷ v[ vlamⁿ (length fcs) f1 ]) + +rotK : ℕ → ℕ → Maybe ℕ +rotK k n with n ℕOR.≟ k +... | lt x = just (suc n) +... | eq x = just zero +... | gt x = nothing + +rotVars : ℕ → R.Term → R.Term +rotVars k = replaceVarWithTerm (map-Maybe 𝒗 ∘S rotK k) + +allTrue : List Bool → Bool +allTrue = foldl _and_ true + +any? : List Bool → Bool +any? [] = false +any? (false ∷ x₁) = any? x₁ +any? (true ∷ x₁) = true + +mbEquation : NBoundaryTerm → Maybe (R.Term × R.Term) +mbEquation (A , []) = nothing +mbEquation (A , x ∷ xs) = + if (allTrue (join $ L.map (λ t → (L.map (not ∘S flip hasVar t ) ([ predℕ (length xs) ]))) + $ join $ L.map (λ (x , y) → x ∷ [ y ]) xs)) + then just x + else nothing + + +nCubeToEq : NBoundaryTerm → R.TC NBoundaryTerm +nCubeToEq (A , []) = pure $ A , [] +nCubeToEq (A , (f0 , f1) ∷ xs) = (A ,_) <$> do + let f0' = ToTerm.toTerm {Unit} {Unit} (defaultCtx (length xs)) + (hco (join $ L.map (λ (k , (x , y)) → + ((rev $ L.insertAt k (just false) (repeat (predℕ (length xs)) nothing)) + , cell (rotVars (predℕ $ length xs) $ x)) + ∷ [ ((rev $ L.insertAt k (just true) (repeat (predℕ (length xs)) nothing)) + , cell (rotVars (predℕ $ length xs) $ y)) ]) + (L.zipWithIndex xs)) + (cell f0)) + + addNDimsToCtx (length xs) $ ⦇ pure (f0' , f1 ) ∷ + (mapM (λ (x , y) → + ⦇ R.reduce (subfaceCellNoDrop (repeat (predℕ $ length xs) (nothing) ∷ʳ just true) x ) + , R.reduce (subfaceCellNoDrop (repeat (predℕ $ length xs) (nothing) ∷ʳ just true) y ) + ⦈) + xs) ⦈ + +nCubeToEqPath : NBoundaryTerm → R.Term +nCubeToEqPath (A , []) = q[ refl ] +nCubeToEqPath (A , (f0 , f1) ∷ xs) = + let dim = (length xs) + in vlamⁿ (suc dim) (ToTerm.toTermFill' {Unit} {Unit} (defaultCtx dim) + (hcodata (join $ L.map (λ (k , (x , y)) → + ((rev $ L.insertAt k (just false) (repeat (predℕ (length xs)) nothing)) + , cell (rotVars (predℕ $ length xs) $ x)) + ∷ [ ((rev $ L.insertAt k (just true) (repeat (predℕ (length xs)) nothing)) + , cell (rotVars (predℕ $ length xs) $ y)) ]) + (L.zipWithIndex xs)) + (cell f0))) + +faceSubFace : ℕ → (Bool × ℕ) → SubFace +faceSubFace dim (b , k) = + take k (repeat (predℕ dim) nothing) + ++ [ just b ] ++ + drop k (repeat (predℕ dim) nothing) + + + +tryCastAsNoCongM : CuTerm → R.TC (CuTerm' ⊥ Unit) +tryCastAsNoCongM = fromJust [ "failed to cast as no cong" ]ₑ ∘S tryCastAsNoCong + + +quoteCuTermNC : Maybe R.Type → ℕ → R.Term → R.TC CuTermNC +quoteCuTermNC mbty dim = quoteCuTerm mbty dim >=> tryCastAsNoCongM + +extractCuTermNC : Maybe R.Type → ℕ → R.Term → R.TC CuTermNC +extractCuTermNC mbTy dim = + (ECT.extractCuTerm' mbTy 100 dim ∘S appNDimsI dim ∘S liftVarsFrom dim 0) + >=> tryCastAsNoCongM + + +quoteBdNC : NBoundaryTerm → R.TC (CuBoundary' ⊥ Unit) +quoteBdNC (A , xs) = do + let dim = predℕ (length xs) + mapM (λ (t0 , t1) → ⦇ quoteCuTermNC (just A) dim t0 , quoteCuTermNC (just A) dim t1 ⦈ ) xs + + + + +extractCuTermFromNPath : R.Type → R.Term → R.TC (ℕ × CuTerm) +extractCuTermFromNPath ty tm = do + (A , dim) ← map-snd length <$> matchNCube ty + (dim ,_) <$> ECT.extractCuTerm' + (just A) + 100 dim (pathAppN dim tm) + + + +macro + showCuTerm : R.Term → R.Term → R.TC Unit + showCuTerm t h = R.withReduceDefs (false , [ quote ua ]) do + hTy ← R.inferType t >>= wait-for-term >>= R.normalise + (dim , cu) ← extractCuTermFromNPath hTy t + te ← ppCT dim 100 cu + R.typeError te + + showCuCode : R.Term → R.Term → R.TC Unit + showCuCode t h = R.withReduceDefs (false , [ quote ua ]) do + hTy ← R.inferType t >>= wait-for-term >>= R.normalise + (dim , cu) ← extractCuTermFromNPath hTy t + c ← codeGen false dim 100 cu + R.typeError [ c ]ₑ diff --git a/Cubical/Tactics/Reflection/Tests/QuoteCubical.agda b/Cubical/Tactics/Reflection/Tests/QuoteCubical.agda new file mode 100644 index 0000000000..7c007d88e1 --- /dev/null +++ b/Cubical/Tactics/Reflection/Tests/QuoteCubical.agda @@ -0,0 +1,201 @@ +{- +This module defines several tests for verifying the functionality of cubical term reflection and extraction provided by `Cubical.Tactics.Reflection.QuoteCubical`. +-} + +{-# OPTIONS --safe -v 3 #-} +module Cubical.Tactics.Reflection.Tests.QuoteCubical where + +import Agda.Builtin.Reflection as R + +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.List as L + +open import Cubical.Reflection.Base hiding (v) + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.Reflection.Utilities + + +open import Cubical.Tactics.Reflection.Dimensions +open import Cubical.Tactics.Reflection.QuoteCubical +open import Cubical.Tactics.Reflection.CuTerm +open import Cubical.Tactics.Reflection.Error + + + + +macro + extractIExprTest : R.Term → R.Term → R.TC Unit + extractIExprTest t h = assertNoErr h do + t' ← extractIExprM t + R.unify t (IExpr→Term t') + + normIExprTest : R.Term → R.Term → R.TC Unit + normIExprTest t h = do + t' ← R.normalise t >>= extractIExprM + R.unify t (IExpr→Term t') + R.unify t (IExpr→Term (normIExpr t')) + wrapError h ("t : " ∷ₑ IExpr→Term t' ∷nl + "norm : " ∷ₑ (IExpr→Term (normIExpr t')) ∷ₑ []) + +_ : ResultIs ✓-pass +_ = extractIExprTest i0 + +_ : ResultIs ✓-pass +_ = extractIExprTest i1 + + + +module _ (i j k l : I) where + _ : ResultIs ✓-pass + _ = extractIExprTest (~ (j ∧ i ∧ ~ j) ∨ k ∨ (i ∧ ~ i) ∨ (l ∧ i)) + + + _ : ResultIs ✓-pass + _ = extractIExprTest (i ∨ ~ i ∨ j ∨ ~ j ∨ k ∨ ~ k) + + _ : ResultIs ✓-pass + _ = extractIExprTest (~ (i ∨ ~ i ∨ j ∨ ~ j ∨ k ∨ ~ k)) + + _ : ResultIs + ("t : i ∨ i ∨ (i ∧ i ∧ j) ∨ i ∧ i ∧ k " ∷ + "norm : i " ∷ []) + _ = normIExprTest (i ∨ i ∨ (i ∧ i ∧ (j ∨ k))) + + _ : ResultIs + ("t : (i ∧ k) ∨ j ∧ k " ∷ + "norm : (j ∧ k) ∨ i ∧ k " ∷ []) + _ = normIExprTest ((i ∧ k) ∨ (j ∧ k)) + + + +module _ (f : ℕ → I) (i j k l : ℕ) where + + _ : ResultIs ⊘-fail + _ = extractIExprTest (~ (f j ∧ f i ∧ ~ f j) ∨ (f k) ∨ (f i ∧ ~ f i) ∨ (f l ∧ f i)) + + +macro + + extarctCuTermTest : R.Term → R.Term → R.TC Unit + extarctCuTermTest t h = assertNoErr h do + hTy ← R.inferType t >>= wait-for-term >>= R.normalise + (dim , cu) ← extractCuTermFromNPath hTy t + (R.unify t (toTerm dim cu)) + + + +module gencode (A : Type) (x y z w v : A) + (p : x ≡ y)(q : y ≡ z)(r : z ≡ w) (s : w ≡ v) where + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (assoc p q r) + + -- _ : ResultIs ✓-pass + -- _ = extarctCuTermTest (pentagonIdentity p q r s) + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (doubleCompPath-filler p q r) + + + +module _ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) + (p : x ≡ y) + {u v : A} (q : u ≡ v) where + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (cong₂Funct f p q) + + + +module E5 (A B C D E F : Type) + (e₀ : A ≃ B) (e₁ : B ≃ C) (e₂ : C ≃ D) (e₃ : D ≃ E) (e₄ : E ≃ F) where + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (ua e₀ ∙∙ ua e₁ ∙∙ ua e₂ ∙ ua e₃ ∙ ua e₄) + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (cong₂Funct (λ A B → List (A × B)) (ua e₁) (ua e₂)) + + +module _ where + open import Cubical.HITs.S2 + open import Cubical.HITs.Susp + + ss : SquareP (λ u v → Square {A = Susp S²} _ _ _ _) _ _ _ _ + ss i j u v = toSuspPresInvS² (surf u v) i j + + _ : ResultIs ⊘-fail + _ = extarctCuTermTest ss + + +module _ where + open import Cubical.HITs.S1 + open import Cubical.Data.Int + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (decodeSquare (pos 4)) + +module _ {ℓ} {A : Type ℓ} {x y z w : A} + {B C : Type ℓ} + (f f' : A → B) + (g : B → B → C) + (p p' : Path A x y) (q q' : y ≡ z) (r : z ≡ w) where + + + pp = cong₂ g (cong f (p ∙' q)) (cong f' (p' ∙ q')) + + _ : ResultIs ✓-pass + _ = extarctCuTermTest pp + + _ : ResultIs ✓-pass + _ = extarctCuTermTest (cong-∙ f p q) + + + + +module _ (A B : Type) {x y z : A} (f : A → A → B) + (p : _≡_ {A = (A → B)} (λ x' → f x' x) (λ x' → f x' y)) + (q : _≡_ {A = (A → B)} (λ x' → f x' y) (λ x' → f x' z)) + where + + _ : ResultIs ✓-pass + _ = extarctCuTermTest ((p ∙∙ q ∙∙ (sym q))) + + +module _ (A B : Type) {x y z : A} (f : A → A → B) + (p : _≡_ {A = ({A} → B)} (λ {x'} → f x' x) (λ {x'} → f x' y)) + (q : _≡_ {A = ({A} → B)} (λ {x'} → f x' y) (λ {x'} → f x' z)) where + + + _ : ResultIs ⊘-fail + _ = extarctCuTermTest (p ∙∙ q ∙∙ sym q) + +module _ (A B : Type) {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : x ≡ z) where + + w w' : Path (B → A) (λ _ → x) λ _ → z + w i = λ _ → (p ∙ q) i + w' = cong const p ∙ cong const q + + w≡w'' : w ≡ w' + w≡w'' = refl + + + + _ : ResultIs ⊘-fail + _ = extarctCuTermTest w + + _ : ResultIs ✓-pass + _ = extarctCuTermTest w' + diff --git a/Cubical/Tactics/Reflection/Utilities.agda b/Cubical/Tactics/Reflection/Utilities.agda index a34fc71b0d..5fc7fbb07d 100644 --- a/Cubical/Tactics/Reflection/Utilities.agda +++ b/Cubical/Tactics/Reflection/Utilities.agda @@ -1,17 +1,83 @@ +{- + +This module provides utilities for traversing the reflected representation of terms using monadic operations. It is designed to apply operations specifically at `variable`, `constructor`, and `definition` nodes. It also handles the lifting of terms inside lambda expressions and pattern lambdas. + +### General Overview + +- **Functionality**: + - Assists in traversing reflected terms. + - Applies operations at specific nodes (`variable`, `constructor`, `definition`). + - Manages lifting within lambda and pattern lambda expressions. + +- **Usages**: + - Can act either as a fold or a map depending on the application with diferent monads. + - Facilitates easier refactoring of code when the reflected representation changes. + - Handles termination issues that are cumbersome when functions traverse or map over reflected terms. + - Allows for the easy inclusion of state or error handling while refactoring functions operating over terms. + +#### Specialized Transformations + +- **`remapVars`**: + - Remaps variables using a specified function. + +- **`replaceVarWithCon`**: + - Replaces variables with constructors based on a mapping function. + +- **`liftVars`** and **`liftVarsFrom`**: + - Functions to lift variables by a specified amount. + +- **`dropVar`**: + - Drops a specified variable. + +#### Modules for Specific Transformations + +- **`LiftFrom`**: + - Provides a public interface for lifting variables by a specified amount. + +- **`dropVars`**: + - Offers a function to drop a specified number of variables. + +#### Additional Utilities + +- **`invVar`**: + - Replaces Interval variable with a negated form. + +- **`hasVar`** and **`hasVarBy`**: + - Checks for the presence of a variable based on given conditions. + +- **`findInterval`**: + - Finds an interval within a term. + +- **`replaceVarWithTerm`**: + - Replaces variables with terms based on a mapping function. + +- **`substTms`**: + - Substitutes a list of terms into a given term. + +- **`replaceAtTrm`**: + - Replaces a term at a specified variable position. + +-} + module Cubical.Tactics.Reflection.Utilities where open import Cubical.Foundations.Prelude hiding (Type) +open import Cubical.Foundations.Function open import Agda.Builtin.Reflection hiding (Type) open import Agda.Builtin.String -open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_) +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_) open import Cubical.Reflection.Base -open import Cubical.Data.List -open import Cubical.Data.Maybe +open import Cubical.Reflection.Sugar +open import Cubical.Data.List as L +import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Maybe as Mb +open import Cubical.Data.Bool +open import Cubical.Data.Unit open import Cubical.Data.FinData using () renaming (zero to fzero; suc to fsuc) open import Cubical.Data.Sigma -open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Nat using (ℕ; suc; _+_; zero; _∸_; discreteℕ; predℕ) open import Cubical.Tactics.Reflection open import Cubical.Tactics.Reflection.Variables @@ -32,3 +98,316 @@ errorOut' something = typeError (strErr "Don't know what to do with " ∷ termEr _==_ = primQNameEquality {-# INLINE _==_ #-} + +mapArg : ∀ {ℓ ℓ'} → {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) → Arg A → Arg B +mapArg f (arg i x) = arg i (f x) + +unArg : ∀ {ℓ} → {A : Type ℓ} → Arg A → A +unArg (arg i x) = x + +argInfo : ∀ {ℓ} → {A : Type ℓ} → Arg A → ArgInfo +argInfo (arg i x) = i + +module atVarOrConOrDefMmp {M : Functorω} + {{RA : RawApplicative M}} {{_ : RawMonad M {{RA}} }} + (f : ℕ → ℕ → (List (Arg Term)) → M (List (Arg Term)) → (List (M (Arg Term))) → (M Term)) + (h : ℕ → Name → (List (Arg Term)) → M (List (Arg Term)) → (List (M (Arg Term))) → (M Term)) + (g : ℕ → Name → (List (Arg Term)) → M (List (Arg Term)) → (List (M (Arg Term))) → (M Term)) + where + + ra : ℕ → List (Arg Term) → M (List (Arg Term)) + raT : ℕ → List (Arg Term) → (List (M (Arg Term))) + rc : ℕ → List Clause → M (List Clause) + rcl : ℕ → Clause → M Clause + rtl : ℕ → Telescope → M Telescope + rv : ℕ → Term → M Term + + rpt : ℕ → Pattern → M Pattern + rpts : ℕ → List (Arg Pattern) → M (List (Arg Pattern)) + rs : ℕ → Sort → M Sort + + ra n [] = ⦇ [] ⦈ + ra n (arg i x ∷ x₂) = + ⦇ ⦇ (arg i) (rv n x) ⦈ ∷ ra n x₂ ⦈ + + raT n [] = [] + raT n (arg i x ∷ x₂) = + ⦇ (arg i) (rv n x) ⦈ ∷ raT n x₂ + + rv n (var x args) = f n x args (ra n args) (raT n args) + rv n (con c args) = h n c args (ra n args) (raT n args) + rv n (def f' args) = g n f' args (ra n args) (raT n args) + rv n (lam v₁ (abs s x)) = + (lam v₁) <$> (abs s <$> (rv (suc n) x)) + rv n (pat-lam cs args) = ⦇ pat-lam (rc n cs) (ra n args) ⦈ + rv n (pi (arg i x) (abs s x₁)) = + ⦇ pi ((arg i) <$> rv n x) (abs s <$> rv (suc n) x₁) ⦈ + rv n (agda-sort s) = ⦇ agda-sort (rs n s) ⦈ + rv n (lit l) = ⦇ (lit l) ⦈ + rv n (meta x x₁) = ⦇ meta ⦇ x ⦈ (ra n x₁) ⦈ + rv n unknown = ⦇ unknown ⦈ + + + rs n (set t) = ⦇ set (rv n t) ⦈ + rs n (lit n₁) = ⦇ (lit n₁) ⦈ + rs n (prop t) = ⦇ prop (rv n t) ⦈ + rs n (propLit n₁) = ⦇ (propLit n₁) ⦈ + rs n (inf n₁) = ⦇ (inf n₁) ⦈ + rs n unknown = ⦇ unknown ⦈ + + rtl n [] = ⦇ [] ⦈ + rtl n ((v , arg i x) ∷ xs) = + ⦇ ⦇ ⦇ v ⦈ , ⦇ arg ⦇ i ⦈ (rv n x) ⦈ ⦈ ∷ rtl (suc n) xs ⦈ + + rc n [] = ⦇ [] ⦈ + rc n (cl ∷ cls) = + ⦇ rcl n cl ∷ rc n cls ⦈ + + + rcl n (clause tel ps t) = + ⦇ clause (rtl n tel) (rpts n ps) (rv (length tel + n) t) ⦈ + rcl n (absurd-clause tel ps) = + ⦇ absurd-clause (rtl n tel) (rpts n ps) ⦈ + + + + rpt n (con c ps) = con c <$> (rpts n ps) + rpt n (dot t) = dot <$> (rv n t) + rpt n (var x) = pure $ var x + rpt n (lit l) = pure $ lit l + rpt n (proj f) = pure $ proj f + rpt n (absurd x) = pure $ absurd x + + rpts n [] = ⦇ [] ⦈ + rpts n ((arg i x) ∷ xs) = ⦇ ⦇ arg ⦇ i ⦈ (rpt n x) ⦈ ∷ rpts n xs ⦈ + + rv0 = rv 0 + +atVarOrConOrDefMmp = atVarOrConOrDefMmp.rv0 + + +module atVarOrDefMmp {M : Functorω} + {{RA : RawApplicative M}} {{RM : RawMonad M {{RA}} }} + (f : ℕ → ℕ → (List (Arg Term)) → M (List (Arg Term)) → (List (M (Arg Term))) → (M Term)) + (g : ℕ → Name → (List (Arg Term)) → M (List (Arg Term)) → (List (M (Arg Term))) → (M Term)) + where + open atVarOrConOrDefMmp {M} {{RA}} {{RM}} + f + (λ n c _ argsM _ → _<$>_ {M = M} (con c) argsM) + g public + + +atVarOrDefMmp = atVarOrDefMmp.rv0 + + +module atVarOrDefM {M : Functorω} + {{RA : RawApplicative M}} {{RM : RawMonad M {{RA}} }} + (f : ℕ → ℕ → (List (Arg Term)) → M (List (Arg Term)) → (M Term)) + (g : ℕ → Name → (List (Arg Term)) → M (List (Arg Term)) → (M Term)) + where + + open atVarOrDefMmp {M = M} + {{RA}} {{RM }} + (λ n k l l' _ → f n k l l') + (λ n k l l' _ → g n k l l') public + +atVarOrDefM = atVarOrDefM.rv0 + +atVarOrM : (ℕ → ℕ → List (Arg Term) → Maybe Term) → (ℕ → Name → List (Arg Term) → Maybe Term) → Term → Term +atVarOrM f g = rv zero + where + open atVarOrDefM {{_}} {{RawMonadIdentityM}} + (λ n k _ args → + let t = var k args + t' = (Mb.fromJust-def t (f n (k ∸ n) args)) + in (if (k <ℕ n) then t else t')) + (λ n nm _ args → + let t = def nm args + in Mb.fromJust-def t (g n nm args)) + +atVarOrM' : (ℕ → ℕ → List (Arg Term) → Maybe Term) → (ℕ → Name → List (Arg Term) → Maybe Term) → Term → Term +atVarOrM' f g = rv zero + where + open atVarOrDefM {{_}} {{RawMonadIdentityM}} + (λ n k args0 args → + let t = var k args + t' = (Mb.fromJust-def t (f n (k ∸ n) args0)) + in (if (k <ℕ n) then t else t')) + (λ n nm args0 args → + Mb.fromJust-def (def nm args) (g n nm args0)) + +atVarOrConM' : (ℕ → ℕ → List (Arg Term) → Maybe Term) → + (ℕ → Name → List (Arg Term) → Maybe Term) + → (ℕ → Name → List (Arg Term) → Maybe Term) → Term → Term +atVarOrConM' f h g = rv zero + where + open atVarOrConOrDefMmp {{_}} {{RawMonadIdentityM}} + (λ n k args0 args _ → + let t = var k args + t' = (Mb.fromJust-def t (f n (k ∸ n) args0)) + in (if (k <ℕ n) then t else t')) + (λ n nm args0 args _ → + Mb.fromJust-def (con nm args) (h n nm args0)) + (λ n nm args0 args _ → + Mb.fromJust-def (def nm args) (g n nm args0)) + + + +module atVarM {M : Functorω} + {{RA : RawApplicative M}} {{RM : RawMonad M {{RA}} }} + (f : ℕ → ℕ → List (Arg Term) → Maybe (M Term)) where + + + open atVarOrDefM + (λ n k _ args → RawMonad._>>=_ RM args λ args → + let t = var k args + in (Mb.fromJust-def (RawApplicative.pure RA t) (if (k <ℕ n) then nothing else (f n (k ∸ n) args)))) + (λ n nm _ args → RawMonad._>>=_ RM args λ args → RawApplicative.pure RA (def nm args)) + public + +module atVar (f : ℕ → ℕ → List (Arg Term) → Maybe (Term)) where + + open atVarM + {{_}} + {{RawMonadIdentityM}} f + public + +atVarM : {M : Functorω} + {{RA : RawApplicative M}} {{_ : RawMonad M {{RA}} }} + (f : ℕ → ℕ → List (Arg Term) → Maybe (M Term)) → Term → M Term +atVarM f = atVarM.rv f zero + + +atVar : (ℕ → ℕ → List (Arg Term) → Maybe Term) → Term → Term +atVar f = atVar.rv f zero + +remapVars : (ℕ → ℕ) → Term → Term +remapVars f = atVar λ n k args → just (var (f k + n) args) + + +-- only for not applied vars!! +replaceVarWithCon : (ℕ → (Maybe Name)) → Term → Term +replaceVarWithCon f = atVar λ n k args → map-Maybe (λ nm → con nm []) (f k) + +liftVars : Term → Term +liftVars = atVar λ n k args → just (var (n + suc k) args) + +liftVarsFrom : ℕ → ℕ → Term → Term +liftVarsFrom m = atVar.rv (λ n k args → just (var (n + m + k) args)) + + +module LiftFrom (m : ℕ) where + open atVar (λ n k args → just (var (n + m + k) args)) public + + + +dropVar : ℕ → Term → Term +dropVar = atVar.rv (λ n k args → just (var (n + predℕ k) args)) + + + +module dropVars (m : ℕ) where + open atVar (λ n k args → just (var (n + (k ∸ m)) args)) public + + + +invVar : ℕ → Term → Term +invVar m = atVar λ where + n k [] → if (Dec→Bool (discreteℕ m k) ) + then just (def (quote ~_) v[ var (k + n) [] ]) + else nothing + _ _ _ → nothing + + + + + + +hasVar : ℕ → Term → Bool +hasVar k' tm = snd $ runIdentity $ (unwrap (atVarM f tm) false) + where + f : ℕ → ℕ → List (Arg Term) → Maybe ([ State₀T Bool RMT IdentityF ] Term) + f n k args = just (wrap (pure ∘S ((var (n + k) args ,_)) ∘S (λ where + true → true + false → k =ℕ k' + ))) + +hasVarBy : (ℕ → Bool) → Term → Bool +hasVarBy g tm = snd $ runIdentity $ (unwrap (atVarM f tm) false) + where + f : ℕ → ℕ → List (Arg Term) → Maybe ([ State₀T Bool RMT IdentityF ] Term) + f n k args = just (wrap (pure ∘S ((var (n + k) args ,_)) ∘S (λ where + true → true + false → g k + ))) + + + +findInterval : ℕ → Term → Maybe Term +findInterval dim tm = + snd $ runIdentity $ (unwrap {T = State₀T (Maybe Term)} (atVarOrDefM.rv + (λ n k _ args' → + let z = (runIdentity (unwrap args' nothing)) + in wrap (identity ∘S (var k (fst z) ,_) ∘S + (_mb>> snd z ∘S _mb>> f n k (fst z)))) + (λ n nm _ args' → + let z = (runIdentity (unwrap args' nothing)) + in wrap + (identity ∘S (def nm (fst z) ,_) ∘S + (_mb>> snd z + ∘S _mb>> (map-Maybe (dropVars.rv n zero) (g nm (fst z) (def nm (fst z))))))) + 0 tm) nothing) + where + + _mb>>_ : Maybe Term → Maybe Term → Maybe Term + nothing mb>> x₁ = x₁ + just x mb>> _ = just x + + + f : ℕ → ℕ → List (Arg Term) → Maybe Term + f n x [] = + if (n <ℕ (suc x)) and (x <ℕ (n + dim)) + then just (var (x ∸ n) []) + else nothing + f n k (x ∷ args) = nothing + + g : Name → List (Arg Term) → Term → Maybe Term + g (quote _∨_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote _∧_) a@(_ v∷ v[ _ ]) tm = just tm + g (quote ~_) a@(v[ _ ]) tm = just tm + g _ _ _ = nothing + + +replaceVarWithTerm : (ℕ → Maybe Term) → Term → Term +replaceVarWithTerm f = + atVar λ n k _ → + map-Maybe (liftVarsFrom n zero) (f k) + + +substTms : List Term → Term → Term +substTms xs = dropVars.rv (length xs) zero ∘S replaceVarWithTerm (lookupMb (map (liftVarsFrom (length xs) 0) xs)) + +replaceAtTrm : ℕ → Term → Term → Term +replaceAtTrm k t = + dropVar k ∘ replaceVarWithTerm (z k) + + where + z : ℕ → ℕ → Maybe Term + z zero zero = just t + z zero (suc y) = nothing + z (suc x) zero = nothing + z (suc x) (suc y) = z x y + + +fromJust : ∀ {ℓ} {A : Type ℓ} → List ErrorPart → Maybe A → TC A +fromJust e = Mb.rec (typeError e) pure + +liftedTele : Telescope → Telescope +liftedTele [] = [] +liftedTele (x ∷ xs) = L.map (map-snd (mapArg liftVars)) (x ∷ liftedTele xs) + +macro + q[_] : Term → Term → TC Unit + q[_] tm h = + quoteTC tm >>= quoteTC >=> unify h diff --git a/Cubical/Tactics/Reflection/Variables.agda b/Cubical/Tactics/Reflection/Variables.agda index 66952ccd25..c0c42325b6 100644 --- a/Cubical/Tactics/Reflection/Variables.agda +++ b/Cubical/Tactics/Reflection/Variables.agda @@ -9,18 +9,20 @@ module Cubical.Tactics.Reflection.Variables where open import Cubical.Foundations.Prelude hiding (Type) +open import Cubical.Foundations.Function open import Agda.Builtin.Reflection hiding (Type) open import Agda.Builtin.String open import Agda.Builtin.Float open import Agda.Builtin.Word open import Agda.Builtin.Char -open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_) +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_ ; _<_ to _<ℕ_ ; _*_ to _*ℕ_ ; _+_ to _+ℕ_) open import Cubical.Reflection.Base open import Cubical.Data.Bool open import Cubical.Data.List open import Cubical.Data.Maybe +open import Cubical.Data.Sigma open import Cubical.Data.Nat using (ℕ) open import Cubical.Tactics.Reflection @@ -78,3 +80,43 @@ indexOf t (t' ∷ l) = then just 0 else map-Maybe (λ k → ℕ.suc k) (indexOf t l) indexOf t [] = nothing + +infixr 40 _<>_ + +_<>_ = primStringAppend + + +digitsToSubscripts : Char → Char +digitsToSubscripts = λ where + '0' → '₀' ; '1' → '₁' ; '2' → '₂' ; '3' → '₃' ; '4' → '₄' ; '5' → '₅' + '6' → '₆' ; '7' → '₇' ; '8' → '₈' ; '9' → '₉' ; x → x + +subscriptToℕ : Char → Maybe ℕ +subscriptToℕ = λ where + '₀' → just 0 ; '₁' → just 1 ; '₂' → just 2 ; '₃' → just 3 ; '₄' → just 5 ; '₅' → just 5 + '₆' → just 6 ; '₇' → just 7 ; '₈' → just 8 ; '₉' → just 9 ; x → nothing + + + + +getSubscript : String → Maybe (String × ℕ) +getSubscript s = + let s' = rev (primStringToList s) + sbs = takeWhile subscriptToℕ s' + v = length sbs + in if (0 <ℕ v) then + just (primStringFromList (rev (drop (length sbs) s')) , fromBase10rev sbs) + else nothing + where + fromBase10rev : List ℕ → ℕ + fromBase10rev [] = ℕ.zero + fromBase10rev (x ∷ xs) = x +ℕ (10 *ℕ fromBase10rev xs) + +mkNiceVar' : String → ℕ → String +mkNiceVar' v k = v <> + primStringFromList (map digitsToSubscripts (primStringToList (primShowNat k))) + + +mkNiceVar : ℕ → String +mkNiceVar = mkNiceVar' "𝒙" +