diff --git a/lib/Quox/Syntax/Term/Tighten.idr b/lib/Quox/Syntax/Term/Tighten.idr index 24c7018..c954070 100644 --- a/lib/Quox/Syntax/Term/Tighten.idr +++ b/lib/Quox/Syntax/Term/Tighten.idr @@ -314,6 +314,19 @@ dsqueeze (S names (Y body)) = dsqueeze (S names (N body)) = S names $ N body +export +squeezed : {s : Nat} -> ScopeTermN s d n -> + Either (BContext s, Term d (s + n)) (Term d n) +squeezed (S ns (N t)) = Right t +squeezed (S ns (Y t)) = maybe (Left (ns, t)) Right $ tightenN s t + +export +dsqueezed : {s : Nat} -> DScopeTermN s d n -> + Either (BContext s, Term (s + d) n) (Term d n) +dsqueezed (S ns (N t)) = Right t +dsqueezed (S ns (Y t)) = maybe (Left (ns, t)) Right $ tightenN s t @{TermD} + + public export %inline CompH' : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> @@ -324,7 +337,6 @@ CompH' {ty, p, q, val, r, zero, one, loc} = Comp { ty = dsub1 ty q, p, q, val = E $ Coe ty p q val val.loc, r, - -- [fixme] better locations for these vars? zero = DST zero.names $ E $ Coe ty' (B VZ zero.loc) (weakD 1 q) zero.term zero.loc, one = DST one.names $ E $ diff --git a/lib/Quox/Whnf/Main.idr b/lib/Quox/Whnf/Main.idr index 008ed9c..020f1f0 100644 --- a/lib/Quox/Whnf/Main.idr +++ b/lib/Quox/Whnf/Main.idr @@ -90,7 +90,7 @@ CanWhnf Elim Interface.isRedexE where whnf defs ctx $ CaseNat pi piIH (Ann val (dsub1 ty q) val.loc) ret zer suc caseLoc Right nn => pure $ - Element (CaseNat pi piIH nat ret zer suc caseLoc) $ natnf `orNo` nn + Element (CaseNat pi piIH nat ret zer suc caseLoc) (natnf `orNo` nn) -- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝ -- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p] @@ -104,7 +104,7 @@ CanWhnf Elim Interface.isRedexE where Coe ty p q val _ => boxCoe defs ctx pi ty p q val ret body caseLoc Right nb => - pure $ Element (CaseBox pi box ret body caseLoc) $ boxnf `orNo` nb + pure $ Element (CaseBox pi box ret body caseLoc) (boxnf `orNo` nb) -- e : Eq (𝑗 ⇒ A) t u ⊢ e @0 ⇝ t ∷ A‹0/𝑗› -- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A‹1/𝑗› @@ -127,7 +127,7 @@ CanWhnf Elim Interface.isRedexE where whnf defs ctx $ ends (Ann (setLoc appLoc l) ty.zero appLoc) (Ann (setLoc appLoc r) ty.one appLoc) e - B {} => pure $ Element (DApp f p appLoc) $ fnf `orNo` ndlh `orNo` Ah + B {} => pure $ Element (DApp f p appLoc) (fnf `orNo` ndlh `orNo` Ah) -- e ∷ A ⇝ e whnf defs ctx (Ann s a annLoc) = do @@ -136,22 +136,28 @@ CanWhnf Elim Interface.isRedexE where Left _ => let E e = s in pure $ Element e $ noOr2 snf Right ne => do Element a anf <- whnf defs ctx a - pure $ Element (Ann s a annLoc) $ ne `orNo` snf `orNo` anf + pure $ Element (Ann s a annLoc) (ne `orNo` snf `orNo` anf) - whnf defs ctx (Coe (S _ (N ty)) _ _ val coeLoc) = - whnf defs ctx $ Ann val ty coeLoc - whnf defs ctx (Coe (S [< i] (Y ty)) p q val coeLoc) = - case p `decEqv` q of - Yes y => whnf defs ctx $ Ann val (ty // one q) coeLoc - No npq => do - Element ty tynf <- whnf defs (extendDim i ctx) ty - case nchoose $ canPushCoe ty val of - Left y => - pushCoe defs ctx i ty p q val coeLoc - Right npc => - -- [fixme] what about ty.zero = ty.one???? - pure $ Element (Coe (SY [< i] ty) p q val coeLoc) - (tynf `orNo` npc `orNo` notYesNo npq) + whnf defs ctx (Coe ty p q val coeLoc) = + -- 𝑖 ∉ fv(A) + -- ------------------------------- + -- coe (𝑖 ⇒ A) @p @q s ⇝ s ∷ A + -- + -- [fixme] needs a real equality check between ty.zero and ty.one + case dsqueezed ty of + Right ty => whnf defs ctx $ Ann val ty coeLoc + Left ([< i], ty) => + case p `decEqv` q of + -- coe (𝑖 ⇒ A) @p @p s ⇝ (s ∷ A‹p/𝑖›) + Yes _ => whnf defs ctx $ Ann val (ty // one p) coeLoc + No npq => do + Element ty tynf <- whnf defs (extendDim i ctx) ty + case nchoose $ canPushCoe ty val of + Left pc => + pushCoe defs ctx i ty p q val coeLoc + Right npc => + pure $ Element (Coe (SY [< i] ty) p q val coeLoc) + (tynf `orNo` npc `orNo` notYesNo npq) whnf defs ctx (Comp ty p q val r zero one compLoc) = case p `decEqv` q of @@ -173,8 +179,9 @@ CanWhnf Elim Interface.isRedexE where Left y => let Ann ty (TYPE u _) _ = ty in reduceTypeCase defs ctx ty u ret arms def tcLoc - Right nt => pure $ - Element (TypeCase ty ret arms def tcLoc) (tynf `orNo` retnf `orNo` nt) + Right nt => + pure $ Element (TypeCase ty ret arms def tcLoc) + (tynf `orNo` retnf `orNo` nt) whnf defs ctx (CloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' id th el whnf defs ctx (DCloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' th id el