143 lines
4.1 KiB
Idris
143 lines
4.1 KiB
Idris
|
module Quox.Syntax.Var
|
||
|
|
||
|
import Quox.Name
|
||
|
|
||
|
import Data.Nat
|
||
|
import Data.Maybe
|
||
|
import Data.List
|
||
|
import Quox.Pretty
|
||
|
|
||
|
%default total
|
||
|
|
||
|
|
||
|
public export
|
||
|
data Var : Nat -> Type where
|
||
|
VZ : Var (S n)
|
||
|
VS : Var n -> Var (S n)
|
||
|
%name Var i, j
|
||
|
%builtin Natural Var
|
||
|
|
||
|
public export
|
||
|
(.nat) : Var n -> Nat
|
||
|
(VZ).nat = 0
|
||
|
(VS i).nat = S i.nat
|
||
|
%transform "Var.(.nat)" Var.(.nat) i = believe_me i
|
||
|
|
||
|
public export Cast (Var n) Nat where cast n = n.nat
|
||
|
|
||
|
export Eq (Var n) where i == j = i.nat == j.nat
|
||
|
export Ord (Var n) where compare i j = compare i.nat j.nat
|
||
|
export Show (Var n) where showPrec d i = showCon d "V" $ showArg i.nat
|
||
|
|
||
|
|
||
|
private
|
||
|
prettyIndex : Nat -> Pretty.M (Doc a)
|
||
|
prettyIndex i =
|
||
|
ifUnicode (pretty $ pack $ map sup $ unpack $ show i) (":" <+> pretty i)
|
||
|
where
|
||
|
sup : Char -> Char
|
||
|
sup c = case c of
|
||
|
'0' => '⁰'; '1' => '¹'; '2' => '²'; '3' => '³'; '4' => '⁴'
|
||
|
'5' => '⁵'; '6' => '⁶'; '7' => '⁷'; '8' => '⁸'; '9' => '⁹'; _ => c
|
||
|
|
||
|
||| `prettyVar hlok hlerr names i` pretty prints the de Bruijn index `i`.
|
||
|
|||
|
||
|
||| If it is within the bounds of `names`, then it uses the name at that index,
|
||
|
||| highlighted as `hlok`. Otherwise it is just printed as a number highlighted
|
||
|
||| as `hlerr`.
|
||
|
export
|
||
|
prettyVar' : HL -> HL -> List Name -> Nat -> Pretty.M (Doc HL)
|
||
|
prettyVar' hlok hlerr names i =
|
||
|
case inBounds i names of
|
||
|
Yes _ => hlF' hlok [|prettyM (index i names) <+> prettyIndex i|]
|
||
|
No _ => pure $ hl hlerr $ pretty i
|
||
|
|
||
|
export %inline
|
||
|
prettyVar : HL -> HL -> List Name -> Var n -> Pretty.M (Doc HL)
|
||
|
prettyVar hlok hlerr names i = prettyVar' hlok hlerr names i.nat
|
||
|
|
||
|
|
||
|
public export
|
||
|
fromNatWith : (i : Nat) -> (0 p : i `LT` n) -> Var n
|
||
|
fromNatWith Z (LTESucc _) = VZ
|
||
|
fromNatWith (S i) (LTESucc p) = VS $ fromNatWith i p
|
||
|
%transform "Var.fromNatWith" fromNatWith i p = believe_me i
|
||
|
|
||
|
public export %inline
|
||
|
V : (i : Nat) -> {auto 0 p : i `LT` n} -> Var n
|
||
|
V i {p} = fromNatWith i p
|
||
|
|
||
|
public export %inline
|
||
|
tryFromNat : Alternative f => (n : Nat) -> Nat -> f (Var n)
|
||
|
tryFromNat n i =
|
||
|
case i `isLT` n of
|
||
|
Yes p => pure $ fromNatWith i p
|
||
|
No _ => empty
|
||
|
|
||
|
export
|
||
|
0 toNatLT : (i : Var n) -> i.nat `LT` n
|
||
|
toNatLT VZ = LTESucc LTEZero
|
||
|
toNatLT (VS i) = LTESucc $ toNatLT i
|
||
|
|
||
|
export
|
||
|
0 toNatInj : {i, j : Var n} -> i.nat = j.nat -> i = j
|
||
|
toNatInj {i = VZ} {j = VZ} Refl = Refl
|
||
|
toNatInj {i = VZ} {j = (VS i)} Refl impossible
|
||
|
toNatInj {i = (VS i)} {j = VZ} Refl impossible
|
||
|
toNatInj {i = (VS i)} {j = (VS j)} prf =
|
||
|
cong VS $ toNatInj $ succInjective _ _ prf
|
||
|
|
||
|
export
|
||
|
0 fromToNat : (i : Var n) -> (p : i.nat `LT` n) -> fromNatWith i.nat p = i
|
||
|
fromToNat VZ (LTESucc p) = Refl
|
||
|
fromToNat (VS i) (LTESucc p) = rewrite fromToNat i p in Refl
|
||
|
|
||
|
export
|
||
|
0 toFromNat : (i : Nat) -> (p : i `LT` n) -> (fromNatWith i p).nat = i
|
||
|
toFromNat 0 (LTESucc x) = Refl
|
||
|
toFromNat (S k) (LTESucc x) = cong S $ toFromNat k x
|
||
|
|
||
|
|
||
|
public export
|
||
|
(.int) : Var n -> Integer
|
||
|
i.int = natToInteger i.nat
|
||
|
%builtin NaturalToInteger Var.(.int)
|
||
|
|
||
|
public export Cast (Var n) Integer where cast i = i.int
|
||
|
|
||
|
|
||
|
-- not using %transform like other things because weakSpec requires the proof
|
||
|
-- to be relevant. but since only `LTESucc` is ever possible that seems
|
||
|
-- to be a bug?
|
||
|
export
|
||
|
weak : (0 p : m `LTE` n) -> Var m -> Var n
|
||
|
weak p i = fromNatWith i.nat $ transitive (toNatLT i) p {rel=LTE}
|
||
|
|
||
|
public export
|
||
|
0 weakSpec : m `LTE` n -> Var m -> Var n
|
||
|
weakSpec LTEZero _ impossible
|
||
|
weakSpec (LTESucc p) VZ = VZ
|
||
|
weakSpec (LTESucc p) (VS i) = VS $ weakSpec p i
|
||
|
|
||
|
export
|
||
|
0 weakSpecCorrect : (p : m `LTE` n) -> (i : Var m) -> (weakSpec p i).nat = i.nat
|
||
|
weakSpecCorrect LTEZero _ impossible
|
||
|
weakSpecCorrect (LTESucc x) VZ = Refl
|
||
|
weakSpecCorrect (LTESucc x) (VS i) = cong S $ weakSpecCorrect x i
|
||
|
|
||
|
export
|
||
|
0 weakCorrect : (p : m `LTE` n) -> (i : Var m) -> (weak p i).nat = i.nat
|
||
|
weakCorrect LTEZero _ impossible
|
||
|
weakCorrect (LTESucc p) VZ = Refl
|
||
|
weakCorrect (LTESucc p) (VS i) = cong S $ weakCorrect p i
|
||
|
|
||
|
export
|
||
|
0 weakIsSpec : (p : m `LTE` n) -> (i : Var m) -> weak p i = weakSpec p i
|
||
|
weakIsSpec p i = toNatInj $ trans (weakCorrect p i) (sym $ weakSpecCorrect p i)
|
||
|
|
||
|
|
||
|
public export
|
||
|
interface FromVar f where %inline fromVar : Var n -> f n
|
||
|
|
||
|
public export FromVar Var where fromVar = id
|