remove ope stuff too

This commit is contained in:
rhiannon morris 2023-01-08 15:43:54 +01:00
parent 9dbd0b066c
commit 0c1b3a78c3
5 changed files with 47 additions and 47 deletions

View file

@ -1,196 +0,0 @@
||| "order preserving embeddings", for recording a correspondence between
||| a smaller scope and part of a larger one.
module Quox.OPE
import Quox.NatExtra
import Data.Nat
%default total
public export
data OPE : Nat -> Nat -> Type where
Id : OPE n n
Drop : OPE m n -> OPE m (S n)
Keep : OPE m n -> OPE (S m) (S n)
%name OPE p, q
public export %inline Injective Drop where injective Refl = Refl
public export %inline Injective Keep where injective Refl = Refl
public export
opeZero : {n : Nat} -> OPE 0 n
opeZero {n = 0} = Id
opeZero {n = S n} = Drop opeZero
public export
(.) : OPE m n -> OPE n p -> OPE m p
p . Id = p
Id . q = q
p . Drop q = Drop $ p . q
Drop p . Keep q = Drop $ p . q
Keep p . Keep q = Keep $ p . q
public export
toLTE : {m : Nat} -> OPE m n -> m `LTE` n
toLTE Id = reflexive
toLTE (Drop p) = lteSuccRight $ toLTE p
toLTE (Keep p) = LTESucc $ toLTE p
public export
dropInner' : LTE' m n -> OPE m n
dropInner' LTERefl = Id
dropInner' (LTESuccR p) = Drop $ dropInner' $ force p
public export
dropInner : {n : Nat} -> LTE m n -> OPE m n
dropInner = dropInner' . fromLte
public export
dropInnerN : (m : Nat) -> OPE n (m + n)
dropInnerN 0 = Id
dropInnerN (S m) = Drop $ dropInnerN m
public export
interface Tighten t where
tighten : Alternative f => OPE m n -> t n -> f (t m)
parameters {auto _ : Tighten t} {auto _ : Alternative f}
export
tightenInner : {n : Nat} -> m `LTE` n -> t n -> f (t m)
tightenInner = tighten . dropInner
export
tightenN : (m : Nat) -> t (m + n) -> f (t n)
tightenN m = tighten $ dropInnerN m
export
tighten1 : t (S n) -> f (t n)
tighten1 = tightenN 1
-- [todo] can this be done with fancy nats too?
-- with bitmasks sure but that might not be worth the effort
-- [the next day] it probably isn't
-- public export
-- data OPE' : Nat -> Nat -> Type where
-- None : OPE' 0 0
-- Drop : OPE' m n -> OPE' m (S n)
-- Keep : OPE' m n -> OPE' (S m) (S n)
-- %name OPE' q
-- public export %inline
-- drop' : Integer -> Integer
-- drop' n = n * 2
-- public export %inline
-- keep' : Integer -> Integer
-- keep' n = 1 + 2 * n
-- public export
-- data IsOPE : Integer -> (OPE' m n) -> Type where
-- IsNone : 0 `IsOPE` None
-- IsDrop : (0 _ : m `IsOPE` q) -> drop' m `IsOPE` Drop q
-- IsKeep : (0 _ : m `IsOPE` q) -> keep' m `IsOPE` Keep q
-- %name IsOPE p
-- public export
-- record OPE m n where
-- constructor MkOPE
-- value : Integer
-- 0 spec : OPE' m n
-- 0 prf : value `IsOPE` spec
-- 0 pos : So (value >= 0)
-- private
-- 0 idrisPleaseLearnAboutIntegers : {x, y : Integer} -> x = y
-- idrisPleaseLearnAboutIntegers {x, y} = believe_me $ Refl {x}
-- private
-- 0 natIntPlus : (m, n : Nat) ->
-- natToInteger (m + n) = natToInteger m + natToInteger n
-- natIntPlus m n = idrisPleaseLearnAboutIntegers
-- private
-- 0 shiftTwice : (x : Integer) -> (m, n : Nat) ->
-- x `shiftL` (m + n) = (x `shiftL` m) `shiftL` n
-- shiftTwice x m n = idrisPleaseLearnAboutIntegers
-- private
-- 0 shift1 : (x : Integer) -> (x `shiftL` 1) = 2 * x
-- shift1 x = idrisPleaseLearnAboutIntegers
-- private
-- 0 intPlusComm : (x, y : Integer) -> (x + y) = (y + x)
-- intPlusComm x y = idrisPleaseLearnAboutIntegers
-- private
-- 0 intTimes2Minus1 : (x : Integer) -> 2 * x - 1 = 2 * (x - 1) + 1
-- intTimes2Minus1 x = idrisPleaseLearnAboutIntegers
-- private
-- 0 intPosShift : So (x > 0) -> So (x `shiftL` i > 0)
-- intPosShift p = believe_me Oh
-- private
-- 0 intNonnegDec : {x : Integer} -> So (x > 0) -> So (x - 1 >= 0)
-- intNonnegDec p = believe_me Oh
-- private
-- 0 shiftSucc : (x : Integer) -> (n : Nat) ->
-- x `shiftL` S n = 2 * (x `shiftL` n)
-- shiftSucc x n = Calc $
-- |~ x `shiftL` S n
-- ~~ x `shiftL` (n + 1)
-- ...(cong (x `shiftL`) $ sym $ plusCommutative {})
-- ~~ (x `shiftL` n) `shiftL` 1
-- ...(shiftTwice {})
-- ~~ 2 * (x `shiftL` n)
-- ...(shift1 {})
-- private
-- opeIdVal : (n : Nat) -> Integer
-- opeIdVal n = (1 `shiftL` n) - 1
-- private
-- 0 opeIdValSpec : (n : Nat) -> Integer
-- opeIdValSpec 0 = 0
-- opeIdValSpec (S n) = keep' $ opeIdValSpec n
-- private
-- 0 opeIdValOk : (n : Nat) -> opeIdVal n = opeIdValSpec n
-- opeIdValOk 0 = Refl
-- opeIdValOk (S n) = Calc $
-- |~ (1 `shiftL` S n) - 1
-- ~~ 2 * (1 `shiftL` n) - 1 ...(cong (\x => x - 1) $ shiftSucc {})
-- ~~ 2 * (1 `shiftL` n - 1) + 1 ...(intTimes2Minus1 {})
-- ~~ 1 + 2 * (1 `shiftL` n - 1) ...(intPlusComm {})
-- ~~ 1 + 2 * opeIdValSpec n ...(cong (\x => 1 + 2 * x) $ opeIdValOk {})
-- private
-- 0 opeIdSpec : (n : Nat) -> OPE' n n
-- opeIdSpec 0 = None
-- opeIdSpec (S n) = Keep $ opeIdSpec n
-- private
-- 0 opeIdProof' : (n : Nat) -> opeIdValSpec n `IsOPE` opeIdSpec n
-- opeIdProof' 0 = IsNone
-- opeIdProof' (S n) = IsKeep (opeIdProof' n)
-- private
-- 0 opeIdProof : (n : Nat) -> opeIdVal n `IsOPE` opeIdSpec n
-- opeIdProof n = rewrite opeIdValOk n in opeIdProof' n
-- export
-- opeId : {n : Nat} -> OPE n n
-- opeId {n} = MkOPE {prf = opeIdProof n, pos = intNonnegDec $ intPosShift Oh, _}

View file

@ -7,7 +7,7 @@ import public Quox.Syntax.Universe
import public Quox.Syntax.Qty
import public Quox.Syntax.Dim
import public Quox.Name
import public Quox.OPE
-- import public Quox.OPE
import Quox.Pretty

View file

@ -116,46 +116,46 @@ pushSubstsE' : Elim d n -> Elim d n
pushSubstsE' e = (pushSubstsE e).fst
mutual
-- tightening a term/elim also causes substitutions to be pushed through.
-- this is because otherwise a variable in an unused part of the subst
-- would cause it to incorrectly fail
-- mutual
-- -- tightening a term/elim also causes substitutions to be pushed through.
-- -- this is because otherwise a variable in an unused part of the subst
-- -- would cause it to incorrectly fail
export covering
Tighten (Term d) where
tighten p (TYPE l) =
pure $ TYPE l
tighten p (Pi qty x arg res) =
Pi qty x <$> tighten p arg
<*> tighten p res
tighten p (Lam x body) =
Lam x <$> tighten p body
tighten p (E e) =
E <$> tighten p e
tighten p (CloT tm th) =
tighten p $ pushSubstsTWith' id th tm
tighten p (DCloT tm th) =
tighten p $ pushSubstsTWith' th id tm
-- export covering
-- Tighten (Term d) where
-- tighten p (TYPE l) =
-- pure $ TYPE l
-- tighten p (Pi qty x arg res) =
-- Pi qty x <$> tighten p arg
-- <*> tighten p res
-- tighten p (Lam x body) =
-- Lam x <$> tighten p body
-- tighten p (E e) =
-- E <$> tighten p e
-- tighten p (CloT tm th) =
-- tighten p $ pushSubstsTWith' id th tm
-- tighten p (DCloT tm th) =
-- tighten p $ pushSubstsTWith' th id tm
export covering
Tighten (Elim d) where
tighten p (F x) =
pure $ F x
tighten p (B i) =
B <$> tighten p i
tighten p (fun :@ arg) =
[|tighten p fun :@ tighten p arg|]
tighten p (tm :# ty) =
[|tighten p tm :# tighten p ty|]
tighten p (CloE el th) =
tighten p $ pushSubstsEWith' id th el
tighten p (DCloE el th) =
tighten p $ pushSubstsEWith' th id el
-- export covering
-- Tighten (Elim d) where
-- tighten p (F x) =
-- pure $ F x
-- tighten p (B i) =
-- B <$> tighten p i
-- tighten p (fun :@ arg) =
-- [|tighten p fun :@ tighten p arg|]
-- tighten p (tm :# ty) =
-- [|tighten p tm :# tighten p ty|]
-- tighten p (CloE el th) =
-- tighten p $ pushSubstsEWith' id th el
-- tighten p (DCloE el th) =
-- tighten p $ pushSubstsEWith' th id el
export covering
Tighten (ScopeTerm d) where
tighten p (TUsed body) = TUsed <$> tighten (Keep p) body
tighten p (TUnused body) = TUnused <$> tighten p body
-- export covering
-- Tighten (ScopeTerm d) where
-- tighten p (TUsed body) = TUsed <$> tighten (Keep p) body
-- tighten p (TUnused body) = TUnused <$> tighten p body
public export %inline

View file

@ -2,7 +2,7 @@ module Quox.Syntax.Var
import Quox.Name
import Quox.Pretty
import Quox.OPE
-- import Quox.OPE
import Data.Nat
import Data.List
@ -264,10 +264,10 @@ decEqFromBool i j =
public export %inline DecEq (Var n) where decEq = varDecEq
export
Tighten Var where
tighten Id i = pure i
tighten (Drop q) VZ = empty
tighten (Drop q) (VS i) = tighten q i
tighten (Keep q) VZ = pure VZ
tighten (Keep q) (VS i) = VS <$> tighten q i
-- export
-- Tighten Var where
-- tighten Id i = pure i
-- tighten (Drop q) VZ = empty
-- tighten (Drop q) (VS i) = tighten q i
-- tighten (Keep q) VZ = pure VZ
-- tighten (Keep q) (VS i) = VS <$> tighten q i