mirror of
https://github.com/idris-lang/Idris2.git
synced 2025-01-01 08:07:44 +03:00
[ contrib ] Arithmetic on Fin (#830)
Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
This commit is contained in:
parent
48dc732788
commit
583442b359
@ -17,6 +17,14 @@ data Fin : (n : Nat) -> Type where
|
||||
FZ : Fin (S k)
|
||||
FS : Fin k -> Fin (S k)
|
||||
|
||||
||| Cast between Fins with equal indices
|
||||
public export
|
||||
cast : {n : Nat} -> (0 eq : m = n) -> Fin m -> Fin n
|
||||
cast {n = S _} eq FZ = FZ
|
||||
cast {n = Z} eq FZ impossible
|
||||
cast {n = S _} eq (FS k) = FS (cast (succInjective _ _ eq) k)
|
||||
cast {n = Z} eq (FS k) impossible
|
||||
|
||||
export
|
||||
Uninhabited (Fin Z) where
|
||||
uninhabited FZ impossible
|
||||
@ -176,3 +184,99 @@ DecEq (Fin n) where
|
||||
= case decEq f f' of
|
||||
Yes p => Yes $ cong FS p
|
||||
No p => No $ p . fsInjective
|
||||
|
||||
namespace Equality
|
||||
|
||||
||| Pointwise equality of Fins
|
||||
||| It is sometimes complicated to prove equalities on type-changing
|
||||
||| operations on Fins.
|
||||
||| This inductive definition can be used to simplify proof. We can
|
||||
||| recover proofs of equalities by using `homoPointwiseIsEqual`.
|
||||
public export
|
||||
data Pointwise : Fin m -> Fin n -> Type where
|
||||
FZ : Pointwise FZ FZ
|
||||
FS : Pointwise k l -> Pointwise (FS k) (FS l)
|
||||
|
||||
infix 6 ~~~
|
||||
||| Convenient infix notation for the notion of pointwise equality of Fins
|
||||
public export
|
||||
(~~~) : Fin m -> Fin n -> Type
|
||||
(~~~) = Pointwise
|
||||
|
||||
||| Pointwise equality is reflexive
|
||||
export
|
||||
reflexive : {k : Fin m} -> k ~~~ k
|
||||
reflexive {k = FZ} = FZ
|
||||
reflexive {k = FS k} = FS reflexive
|
||||
|
||||
||| Pointwise equality is symmetric
|
||||
export
|
||||
symmetric : k ~~~ l -> l ~~~ k
|
||||
symmetric FZ = FZ
|
||||
symmetric (FS prf) = FS (symmetric prf)
|
||||
|
||||
||| Pointwise equality is transitive
|
||||
export
|
||||
transitive : j ~~~ k -> k ~~~ l -> j ~~~ l
|
||||
transitive FZ FZ = FZ
|
||||
transitive (FS prf) (FS prg) = FS (transitive prf prg)
|
||||
|
||||
||| Pointwise equality is compatible with cast
|
||||
export
|
||||
castEq : {k : Fin m} -> (0 eq : m = n) -> cast eq k ~~~ k
|
||||
castEq {k = FZ} Refl = FZ
|
||||
castEq {k = FS k} Refl = FS (castEq _)
|
||||
|
||||
||| The actual proof used by cast is irrelevant
|
||||
export
|
||||
congCast : {n, q : Nat} -> {k : Fin m} -> {l : Fin p} ->
|
||||
{0 eq1 : m = n} -> {0 eq2 : p = q} ->
|
||||
k ~~~ l ->
|
||||
cast eq1 k ~~~ cast eq2 l
|
||||
congCast eq = transitive (castEq _) (transitive eq $ symmetric $ castEq _)
|
||||
|
||||
||| Last is congruent wrt index equality
|
||||
export
|
||||
congLast : {m, n : Nat} -> (0 _ : m = n) -> last {n=m} ~~~ last {n}
|
||||
congLast {m = Z} {n = Z} eq = reflexive
|
||||
congLast {m = S _} {n = S _} eq = FS (congLast (succInjective _ _ eq))
|
||||
|
||||
export
|
||||
congShift : (m : Nat) -> k ~~~ l -> shift m k ~~~ shift m l
|
||||
congShift Z prf = prf
|
||||
congShift (S m) prf = FS (congShift m prf)
|
||||
|
||||
||| WeakenN is congruent wrt pointwise equality
|
||||
export
|
||||
congWeakenN : k ~~~ l -> weakenN n k ~~~ weakenN n l
|
||||
congWeakenN FZ = FZ
|
||||
congWeakenN (FS prf) = FS (congWeakenN prf)
|
||||
|
||||
||| Pointwise equality is propositional equality on Fins that have the same type
|
||||
export
|
||||
homoPointwiseIsEqual : {0 k, l : Fin m} -> k ~~~ l -> k === l
|
||||
homoPointwiseIsEqual FZ = Refl
|
||||
homoPointwiseIsEqual (FS prf) = cong FS (homoPointwiseIsEqual prf)
|
||||
|
||||
||| Pointwise equality is propositional equality modulo transport on Fins that
|
||||
||| have provably equal types
|
||||
export
|
||||
hetPointwiseIsTransport :
|
||||
{0 k : Fin m} -> {0 l : Fin n} ->
|
||||
(eq : m === n) -> k ~~~ l -> k === rewrite eq in l
|
||||
hetPointwiseIsTransport Refl = homoPointwiseIsEqual
|
||||
|
||||
export
|
||||
finToNatQuotient : k ~~~ l -> finToNat k === finToNat l
|
||||
finToNatQuotient FZ = Refl
|
||||
finToNatQuotient (FS prf) = cong S (finToNatQuotient prf)
|
||||
|
||||
export
|
||||
weakenNeutral : (k : Fin n) -> weaken k ~~~ k
|
||||
weakenNeutral FZ = FZ
|
||||
weakenNeutral (FS k) = FS (weakenNeutral k)
|
||||
|
||||
export
|
||||
weakenNNeutral : (0 m : Nat) -> (k : Fin n) -> weakenN m k ~~~ k
|
||||
weakenNNeutral m FZ = FZ
|
||||
weakenNNeutral m (FS k) = FS (weakenNNeutral m k)
|
||||
|
@ -3,77 +3,130 @@ module Data.Fin.Extra
|
||||
import Data.Fin
|
||||
import Data.Nat
|
||||
|
||||
import Syntax.WithProof
|
||||
import Syntax.PreorderReasoning
|
||||
|
||||
%default total
|
||||
|
||||
||| Proof that an element **n** of Fin **m** , when converted to Nat is smaller than the bound **m**.
|
||||
-------------------------------
|
||||
--- `finToNat`'s properties ---
|
||||
-------------------------------
|
||||
|
||||
||| A Fin's underlying natural number is smaller than the bound
|
||||
export
|
||||
elemSmallerThanBound : (n : Fin m) -> LT (finToNat n) m
|
||||
elemSmallerThanBound FZ = LTESucc LTEZero
|
||||
elemSmallerThanBound (FS x) = LTESucc (elemSmallerThanBound x)
|
||||
|
||||
||| Proof that application of finToNat the last element of Fin **S n** gives **n**.
|
||||
||| Last's underlying natural number is the bound's predecessor
|
||||
export
|
||||
finToNatLastIsBound : {n : Nat} -> finToNat (Fin.last {n}) = n
|
||||
finToNatLastIsBound {n=Z} = Refl
|
||||
finToNatLastIsBound {n=S k} = rewrite finToNatLastIsBound {n=k} in Refl
|
||||
finToNatLastIsBound {n=S k} = cong S finToNatLastIsBound
|
||||
|
||||
||| Proof that an element **n** of Fin **m** , when converted to Nat is smaller than the bound **m**.
|
||||
||| Weaken does not modify the underlying natural number
|
||||
export
|
||||
finToNatWeakenNeutral : {m : Nat} -> {n : Fin m} -> finToNat (weaken n) = finToNat n
|
||||
finToNatWeakenNeutral {n=FZ} = Refl
|
||||
finToNatWeakenNeutral {m=S (S _)} {n=FS _} = cong S finToNatWeakenNeutral
|
||||
finToNatWeakenNeutral : {n : Fin m} -> finToNat (weaken n) = finToNat n
|
||||
finToNatWeakenNeutral = finToNatQuotient (weakenNeutral n)
|
||||
|
||||
-- ||| Proof that it's possible to strengthen a weakened element of Fin **m**.
|
||||
-- export
|
||||
-- strengthenWeakenNeutral : {m : Nat} -> (n : Fin m) -> strengthen (weaken n) = Right n
|
||||
-- strengthenWeakenNeutral {m=S _} FZ = Refl
|
||||
-- strengthenWeakenNeutral {m=S (S _)} (FS k) = rewrite strengthenWeakenNeutral k in Refl
|
||||
||| WeakenN does not modify the underlying natural number
|
||||
export
|
||||
finToNatWeakenNNeutral : (0 m : Nat) -> (k : Fin n) ->
|
||||
finToNat (weakenN m k) = finToNat k
|
||||
finToNatWeakenNNeutral m k = finToNatQuotient (weakenNNeutral m k)
|
||||
|
||||
||| Proof that it's not possible to strengthen the last element of Fin **n**.
|
||||
||| `Shift k` shifts the underlying natural number by `k`.
|
||||
export
|
||||
finToNatShift : (k : Nat) -> (a : Fin n) -> finToNat (shift k a) = k + finToNat a
|
||||
finToNatShift Z a = Refl
|
||||
finToNatShift (S k) a = cong S (finToNatShift k a)
|
||||
|
||||
-------------------------------------------------
|
||||
--- Inversion function and related properties ---
|
||||
-------------------------------------------------
|
||||
|
||||
||| Compute the Fin such that `k + invFin k = n - 1`
|
||||
export
|
||||
invFin : {n : Nat} -> Fin n -> Fin n
|
||||
invFin FZ = last
|
||||
invFin (FS k) = weaken (invFin k)
|
||||
|
||||
||| The inverse of a weakened element is the successor of its inverse
|
||||
export
|
||||
invFinWeakenIsFS : {n : Nat} -> (m : Fin n) -> invFin (weaken m) = FS (invFin m)
|
||||
invFinWeakenIsFS FZ = Refl
|
||||
invFinWeakenIsFS (FS k) = cong weaken (invFinWeakenIsFS k)
|
||||
|
||||
export
|
||||
invFinLastIsFZ : {n : Nat} -> invFin (last {n}) = FZ
|
||||
invFinLastIsFZ {n = Z} = Refl
|
||||
invFinLastIsFZ {n = S n} = cong weaken (invFinLastIsFZ {n})
|
||||
|
||||
||| `invFin` is involutive (i.e. applied twice it is the identity)
|
||||
export
|
||||
invFinInvolutive : {n : Nat} -> (m : Fin n) -> invFin (invFin m) = m
|
||||
invFinInvolutive FZ = invFinLastIsFZ
|
||||
invFinInvolutive (FS k) = Calc $
|
||||
|~ invFin (invFin (FS k))
|
||||
~~ invFin (weaken (invFin k)) ...( Refl )
|
||||
~~ FS (invFin (invFin k)) ...( invFinWeakenIsFS (invFin k) )
|
||||
~~ FS k ...( cong FS (invFinInvolutive k) )
|
||||
|
||||
--------------------------------
|
||||
--- Strengthening properties ---
|
||||
--------------------------------
|
||||
|
||||
||| It's possible to strengthen a weakened element of Fin **m**.
|
||||
export
|
||||
strengthenWeakenIsRight : (n : Fin m) -> strengthen (weaken n) = Right n
|
||||
strengthenWeakenIsRight FZ = Refl
|
||||
strengthenWeakenIsRight (FS k) = rewrite strengthenWeakenIsRight k in Refl
|
||||
|
||||
||| It's not possible to strengthen the last element of Fin **n**.
|
||||
export
|
||||
strengthenLastIsLeft : {n : Nat} -> strengthen (Fin.last {n}) = Left (Fin.last {n})
|
||||
strengthenLastIsLeft {n=Z} = Refl
|
||||
strengthenLastIsLeft {n=S k} = rewrite strengthenLastIsLeft {n=k} in Refl
|
||||
|
||||
||| Enumerate elements of Fin **n** backwards.
|
||||
||| It's possible to strengthen the inverse of a succesor
|
||||
export
|
||||
invFin : {n : Nat} -> Fin n -> Fin n
|
||||
invFin FZ = last
|
||||
invFin {n=S (S _)} (FS k) = weaken (invFin k)
|
||||
strengthenNotLastIsRight : {n : Nat} -> (m : Fin n) -> strengthen (invFin (FS m)) = Right (invFin m)
|
||||
strengthenNotLastIsRight m = strengthenWeakenIsRight (invFin m)
|
||||
|
||||
||| Proof that an inverse of a weakened element of Fin **n** is a successive of an inverse of an original element.
|
||||
export
|
||||
invWeakenIsSucc : {n : Nat} -> (m : Fin n) -> invFin (weaken m) = FS (invFin m)
|
||||
invWeakenIsSucc FZ = Refl
|
||||
invWeakenIsSucc {n=S (S _)} (FS k) = rewrite invWeakenIsSucc k in Refl
|
||||
|
||||
||| Proof that double inversion of Fin **n** gives the original.
|
||||
export
|
||||
doubleInvFinSame : {n : Nat} -> (m : Fin n) -> invFin (invFin m) = m
|
||||
doubleInvFinSame {n=S Z} FZ = Refl
|
||||
doubleInvFinSame {n=S (S k)} FZ = rewrite doubleInvFinSame {n=S k} FZ in Refl
|
||||
doubleInvFinSame {n=S (S _)} (FS x) = trans (invWeakenIsSucc $ invFin x) (cong FS $ doubleInvFinSame x)
|
||||
|
||||
||| Proof that an inverse of the last element of Fin (S **n**) in FZ.
|
||||
export
|
||||
invLastIsFZ : {n : Nat} -> invFin (Fin.last {n}) = FZ
|
||||
invLastIsFZ {n=Z} = Refl
|
||||
invLastIsFZ {n=S k} = rewrite invLastIsFZ {n=k} in Refl
|
||||
|
||||
-- ||| Proof that it's possible to strengthen an inverse of a succesive element of Fin **n**.
|
||||
-- export
|
||||
-- strengthenNotLastIsRight : (m : Fin (S n)) -> strengthen (invFin (FS m)) = Right (invFin m)
|
||||
-- strengthenNotLastIsRight m = strengthenWeakenNeutral (invFin m)
|
||||
--
|
||||
||| Either tightens the bound on a Fin or proves that it's the last.
|
||||
export
|
||||
strengthen' : {n : Nat} -> (m : Fin (S n)) -> Either (m = Fin.last) (m' : Fin n ** finToNat m = finToNat m')
|
||||
strengthen' : {n : Nat} -> (m : Fin (S n)) ->
|
||||
Either (m = Fin.last) (m' : Fin n ** finToNat m = finToNat m')
|
||||
strengthen' {n = Z} FZ = Left Refl
|
||||
strengthen' {n = S k} FZ = Right (FZ ** Refl)
|
||||
strengthen' {n = S k} (FS m) = case strengthen' m of
|
||||
Left eq => Left $ cong FS eq
|
||||
Right (m' ** eq) => Right (FS m' ** cong S eq)
|
||||
|
||||
----------------------------
|
||||
--- Weakening properties ---
|
||||
----------------------------
|
||||
|
||||
export
|
||||
weakenNZeroIdentity : (k : Fin n) -> weakenN 0 k ~~~ k
|
||||
weakenNZeroIdentity FZ = FZ
|
||||
weakenNZeroIdentity (FS k) = FS (weakenNZeroIdentity k)
|
||||
|
||||
export
|
||||
shiftFSLinear : (m : Nat) -> (f : Fin n) -> shift m (FS f) ~~~ FS (shift m f)
|
||||
shiftFSLinear Z f = reflexive
|
||||
shiftFSLinear (S m) f = FS (shiftFSLinear m f)
|
||||
|
||||
export
|
||||
shiftLastIsLast : (m : Nat) -> {n : Nat} ->
|
||||
shift m (Fin.last {n}) ~~~ Fin.last {n=m+n}
|
||||
shiftLastIsLast Z = reflexive
|
||||
shiftLastIsLast (S m) = FS (shiftLastIsLast m)
|
||||
|
||||
-----------------------------------
|
||||
--- Division-related properties ---
|
||||
-----------------------------------
|
||||
|
||||
||| A view of Nat as a quotient of some number and a finite remainder.
|
||||
public export
|
||||
data FractionView : (n : Nat) -> (d : Nat) -> Type where
|
||||
@ -98,6 +151,9 @@ divMod {ok=_} (S n) (S d) =
|
||||
rewrite sym $ plusSuccRightSucc (q * S d) (finToNat r') in
|
||||
cong S $ trans (sym $ cong (plus (q * S d)) eq') eq
|
||||
|
||||
-------------------
|
||||
--- Conversions ---
|
||||
-------------------
|
||||
|
||||
||| Total function to convert a nat to a Fin, given a proof
|
||||
||| that it is less than the bound.
|
||||
@ -113,4 +169,271 @@ natToFinToNat :
|
||||
-> (lte : LT n m)
|
||||
-> finToNat (natToFinLTE n lte) = n
|
||||
natToFinToNat 0 (LTESucc lte) = Refl
|
||||
natToFinToNat (S k) (LTESucc lte) = rewrite natToFinToNat k lte in Refl
|
||||
natToFinToNat (S k) (LTESucc lte) = cong S (natToFinToNat k lte)
|
||||
|
||||
----------------------------------------
|
||||
--- Result-type changing arithmetics ---
|
||||
----------------------------------------
|
||||
|
||||
||| Addition of `Fin`s as bounded naturals.
|
||||
||| The resulting type has the smallest possible bound
|
||||
||| as illustated by the relations with the `last` function.
|
||||
public export
|
||||
(+) : {m, n : Nat} -> Fin m -> Fin (S n) -> Fin (m + n)
|
||||
(+) FZ y = cast (cong S $ plusCommutative n (pred m)) (weakenN (pred m) y)
|
||||
(+) (FS x) y = FS (x + y)
|
||||
|
||||
||| Multiplication of `Fin`s as bounded naturals.
|
||||
||| The resulting type has the smallest possible bound
|
||||
||| as illustated by the relations with the `last` function.
|
||||
public export
|
||||
(*) : {m, n : Nat} -> Fin (S m) -> Fin (S n) -> Fin (S (m * n))
|
||||
(*) FZ _ = FZ
|
||||
(*) {m = S _} (FS x) y = y + x * y
|
||||
|
||||
--- Properties ---
|
||||
|
||||
-- Relation between `+` and `*` and their counterparts on `Nat`
|
||||
|
||||
export
|
||||
finToNatPlusHomo : {m, n : Nat} -> (x : Fin m) -> (y : Fin (S n)) ->
|
||||
finToNat (x + y) = finToNat x + finToNat y
|
||||
finToNatPlusHomo FZ _
|
||||
= finToNatQuotient $ transitive
|
||||
(castEq _)
|
||||
(weakenNNeutral _ _)
|
||||
finToNatPlusHomo (FS x) y = cong S (finToNatPlusHomo x y)
|
||||
|
||||
export
|
||||
finToNatMultHomo : {m, n : Nat} -> (x : Fin (S m)) -> (y : Fin (S n)) ->
|
||||
finToNat (x * y) = finToNat x * finToNat y
|
||||
finToNatMultHomo FZ _ = Refl
|
||||
finToNatMultHomo {m = S _} (FS x) y = Calc $
|
||||
|~ finToNat (FS x * y)
|
||||
~~ finToNat (y + x * y) ...( Refl )
|
||||
~~ finToNat y + finToNat (x * y) ...( finToNatPlusHomo y (x * y) )
|
||||
~~ finToNat y + finToNat x * finToNat y ...( cong (finToNat y +) (finToNatMultHomo x y) )
|
||||
~~ finToNat (FS x) * finToNat y ...( Refl)
|
||||
|
||||
-- Relations to `Fin`'s `last`
|
||||
|
||||
export
|
||||
plusPreservesLast : (m, n : Nat) -> Fin.last {n=m} + Fin.last {n} = Fin.last
|
||||
plusPreservesLast Z n
|
||||
= homoPointwiseIsEqual $ transitive
|
||||
(castEq _)
|
||||
(weakenNNeutral _ _)
|
||||
plusPreservesLast (S m) n = cong FS (plusPreservesLast m n)
|
||||
|
||||
export
|
||||
multPreservesLast : (m, n : Nat) -> Fin.last {n=m} * Fin.last {n} = Fin.last
|
||||
multPreservesLast Z n = Refl
|
||||
multPreservesLast (S m) n = Calc $
|
||||
|~ last + (last * last)
|
||||
~~ last + last ...( cong (last +) (multPreservesLast m n) )
|
||||
~~ last ...( plusPreservesLast n (m * n) )
|
||||
|
||||
-- General addition properties
|
||||
|
||||
export
|
||||
plusSuccRightSucc : {m, n : Nat} -> (left : Fin m) -> (right : Fin (S n)) ->
|
||||
FS (left + right) ~~~ left + FS right
|
||||
plusSuccRightSucc FZ right = FS $ congCast reflexive
|
||||
plusSuccRightSucc (FS left) right = FS $ plusSuccRightSucc left right
|
||||
|
||||
-- Relations to `Fin`-specific `shift` and `weaken`
|
||||
|
||||
export
|
||||
shiftAsPlus : {m, n : Nat} -> (k : Fin (S m)) ->
|
||||
shift n k ~~~ last {n} + k
|
||||
shiftAsPlus {n=Z} k =
|
||||
symmetric $ transitive (castEq _) (weakenNNeutral _ _)
|
||||
shiftAsPlus {n=S n} k = FS (shiftAsPlus k)
|
||||
|
||||
export
|
||||
weakenNAsPlusFZ : {m, n : Nat} -> (k : Fin n) ->
|
||||
weakenN m k = k + the (Fin (S m)) FZ
|
||||
weakenNAsPlusFZ FZ = Refl
|
||||
weakenNAsPlusFZ (FS k) = cong FS (weakenNAsPlusFZ k)
|
||||
|
||||
export
|
||||
weakenNPlusHomo : {0 m, n : Nat} -> (k : Fin p) ->
|
||||
weakenN n (weakenN m k) ~~~ weakenN (m + n) k
|
||||
weakenNPlusHomo FZ = FZ
|
||||
weakenNPlusHomo (FS k) = FS (weakenNPlusHomo k)
|
||||
|
||||
export
|
||||
weakenNOfPlus :
|
||||
{m, n : Nat} -> (k : Fin m) -> (l : Fin (S n)) ->
|
||||
weakenN w (k + l) ~~~ weakenN w k + l
|
||||
weakenNOfPlus FZ l
|
||||
= transitive (congWeakenN (castEq _))
|
||||
$ transitive (weakenNPlusHomo l)
|
||||
$ symmetric (castEq _)
|
||||
weakenNOfPlus (FS k) l = FS (weakenNOfPlus k l)
|
||||
-- General addition properties (continued)
|
||||
|
||||
export
|
||||
plusZeroLeftNeutral : (k : Fin (S n)) -> FZ + k ~~~ k
|
||||
plusZeroLeftNeutral k = transitive (castEq _) (weakenNNeutral _ k)
|
||||
|
||||
export
|
||||
congPlusLeft : {m, n, p : Nat} -> {k : Fin m} -> {l : Fin n} ->
|
||||
(c : Fin (S p)) -> k ~~~ l -> k + c ~~~ l + c
|
||||
congPlusLeft c FZ
|
||||
= transitive (plusZeroLeftNeutral c)
|
||||
(symmetric $ plusZeroLeftNeutral c)
|
||||
congPlusLeft c (FS prf) = FS (congPlusLeft c prf)
|
||||
|
||||
export
|
||||
plusZeroRightNeutral : (k : Fin m) -> k + FZ ~~~ k
|
||||
plusZeroRightNeutral FZ = FZ
|
||||
plusZeroRightNeutral (FS k) = FS (plusZeroRightNeutral k)
|
||||
|
||||
export
|
||||
congPlusRight : {m, n, p : Nat} -> {k : Fin (S n)} -> {l : Fin (S p)} ->
|
||||
(c : Fin m) -> k ~~~ l -> c + k ~~~ c + l
|
||||
congPlusRight c FZ
|
||||
= transitive (plusZeroRightNeutral c)
|
||||
(symmetric $ plusZeroRightNeutral c)
|
||||
congPlusRight {n = S _} {p = S _} c (FS prf)
|
||||
= transitive (symmetric $ plusSuccRightSucc c _)
|
||||
$ transitive (FS $ congPlusRight c prf)
|
||||
(plusSuccRightSucc c _)
|
||||
congPlusRight {p = Z} c (FS prf) impossible
|
||||
|
||||
export
|
||||
plusCommutative : {m, n : Nat} -> (left : Fin (S m)) -> (right : Fin (S n)) ->
|
||||
left + right ~~~ right + left
|
||||
plusCommutative FZ right
|
||||
= transitive (plusZeroLeftNeutral right)
|
||||
(symmetric $ plusZeroRightNeutral right)
|
||||
plusCommutative {m = S _} (FS left) right
|
||||
= transitive (FS (plusCommutative left right))
|
||||
(plusSuccRightSucc right left)
|
||||
|
||||
export
|
||||
plusAssociative :
|
||||
{m, n, p : Nat} ->
|
||||
(left : Fin m) -> (centre : Fin (S n)) -> (right : Fin (S p)) ->
|
||||
left + (centre + right) ~~~ (left + centre) + right
|
||||
plusAssociative FZ centre right
|
||||
= transitive (plusZeroLeftNeutral (centre + right))
|
||||
(congPlusLeft right (symmetric $ plusZeroLeftNeutral centre))
|
||||
plusAssociative (FS left) centre right = FS (plusAssociative left centre right)
|
||||
|
||||
-------------------------------------------------
|
||||
--- Splitting operations and their properties ---
|
||||
-------------------------------------------------
|
||||
|
||||
||| Converts `Fin`s that are used as indexes of parts to an index of a sum.
|
||||
|||
|
||||
||| For example, if you have a `Vect` that is a concatenation of two `Vect`s and
|
||||
||| you have an index either in the first or the second of the original `Vect`s,
|
||||
||| using this function you can get an index in the concatenated one.
|
||||
public export
|
||||
indexSum : {m : Nat} -> Either (Fin m) (Fin n) -> Fin (m + n)
|
||||
indexSum (Left l) = weakenN n l
|
||||
indexSum (Right r) = shift m r
|
||||
|
||||
||| Extracts an index of the first or the second part from the index of a sum.
|
||||
|||
|
||||
||| For example, if you have a `Vect` that is a concatenation of the `Vect`s and
|
||||
||| you have an index of this `Vect`, you have get an index of either left or right
|
||||
||| original `Vect` using this function.
|
||||
public export
|
||||
splitSum : {m : Nat} -> Fin (m + n) -> Either (Fin m) (Fin n)
|
||||
splitSum {m=Z} k = Right k
|
||||
splitSum {m=S m} FZ = Left FZ
|
||||
splitSum {m=S m} (FS k) = mapFst FS $ splitSum k
|
||||
|
||||
||| Calculates the index of a square matrix of size `a * b` by indices of each side.
|
||||
public export
|
||||
indexProd : {n : Nat} -> Fin m -> Fin n -> Fin (m * n)
|
||||
indexProd FZ = weakenN $ mult (pred m) n
|
||||
indexProd (FS k) = shift n . indexProd k
|
||||
|
||||
||| Splits the index of a square matrix of size `a * b` to indices of each side.
|
||||
public export
|
||||
splitProd : {m, n : Nat} -> Fin (m * n) -> (Fin m, Fin n)
|
||||
splitProd {m=S _} p = case splitSum p of
|
||||
Left k => (FZ, k)
|
||||
Right l => mapFst FS $ splitProd l
|
||||
|
||||
--- Properties ---
|
||||
|
||||
export
|
||||
indexSumPreservesLast :
|
||||
(m, n : Nat) -> indexSum {m} (Right $ Fin.last {n}) ~~~ Fin.last {n=m+n}
|
||||
indexSumPreservesLast Z n = reflexive
|
||||
indexSumPreservesLast (S m) n = FS (shiftLastIsLast m)
|
||||
|
||||
export
|
||||
indexProdPreservesLast : (m, n : Nat) ->
|
||||
indexProd (Fin.last {n=m}) (Fin.last {n}) = Fin.last
|
||||
indexProdPreservesLast Z n = homoPointwiseIsEqual
|
||||
$ transitive (weakenNZeroIdentity last)
|
||||
(congLast (sym $ plusZeroRightNeutral n))
|
||||
indexProdPreservesLast (S m) n = Calc $
|
||||
|~ indexProd (last {n=S m}) (last {n})
|
||||
~~ FS (shift n (indexProd last last)) ...( Refl )
|
||||
~~ FS (shift n last) ...( cong (FS . shift n) (indexProdPreservesLast m n ) )
|
||||
~~ last ...( homoPointwiseIsEqual prf )
|
||||
|
||||
where
|
||||
|
||||
prf : shift (S n) (Fin.last {n = n + m * S n}) ~~~ Fin.last {n = n + S (n + m * S n)}
|
||||
prf = transitive (shiftLastIsLast (S n))
|
||||
(congLast (plusSuccRightSucc n (n + m * S n)))
|
||||
|
||||
export
|
||||
splitSumOfWeakenN : (k : Fin m) -> splitSum {m} {n} (weakenN n k) = Left k
|
||||
splitSumOfWeakenN FZ = Refl
|
||||
splitSumOfWeakenN (FS k) = cong (mapFst FS) $ splitSumOfWeakenN k
|
||||
|
||||
export
|
||||
splitSumOfShift : {m : Nat} -> (k : Fin n) -> splitSum {m} {n} (shift m k) = Right k
|
||||
splitSumOfShift {m=Z} k = Refl
|
||||
splitSumOfShift {m=S m} k = cong (mapFst FS) $ splitSumOfShift k
|
||||
|
||||
export
|
||||
splitOfIndexSumInverse : {m : Nat} -> (e : Either (Fin m) (Fin n)) -> splitSum (indexSum e) = e
|
||||
splitOfIndexSumInverse (Left l) = splitSumOfWeakenN l
|
||||
splitOfIndexSumInverse (Right r) = splitSumOfShift r
|
||||
|
||||
export
|
||||
indexOfSplitSumInverse : {m, n : Nat} -> (f : Fin (m + n)) -> indexSum (splitSum {m} {n} f) = f
|
||||
indexOfSplitSumInverse {m=Z} f = Refl
|
||||
indexOfSplitSumInverse {m=S _} FZ = Refl
|
||||
indexOfSplitSumInverse {m=S _} (FS f) with (indexOfSplitSumInverse f)
|
||||
indexOfSplitSumInverse {m=S _} (FS f) | eq with (splitSum f)
|
||||
indexOfSplitSumInverse {m=S _} (FS _) | eq | Left _ = cong FS eq
|
||||
indexOfSplitSumInverse {m=S _} (FS _) | eq | Right _ = cong FS eq
|
||||
|
||||
|
||||
export
|
||||
splitOfIndexProdInverse : {m : Nat} -> (k : Fin m) -> (l : Fin n) ->
|
||||
splitProd (indexProd k l) = (k, l)
|
||||
splitOfIndexProdInverse FZ l
|
||||
= rewrite splitSumOfWeakenN {n = mult (pred m) n} l in Refl
|
||||
splitOfIndexProdInverse (FS k) l
|
||||
= rewrite splitSumOfShift {m=n} $ indexProd k l in
|
||||
cong (mapFst FS) $ splitOfIndexProdInverse k l
|
||||
|
||||
export
|
||||
indexOfSplitProdInverse : {m, n : Nat} -> (f : Fin (m * n)) ->
|
||||
uncurry (indexProd {m} {n}) (splitProd {m} {n} f) = f
|
||||
indexOfSplitProdInverse {m = S _} f with (@@ splitSum f)
|
||||
indexOfSplitProdInverse {m = S _} f | (Left l ** eq) = rewrite eq in Calc $
|
||||
|~ indexSum (Left l)
|
||||
~~ indexSum (splitSum f) ...( cong indexSum (sym eq) )
|
||||
~~ f ...( indexOfSplitSumInverse f )
|
||||
indexOfSplitProdInverse f | (Right r ** eq) with (@@ splitProd r)
|
||||
indexOfSplitProdInverse f | (Right r ** eq) | ((p, q) ** eq2)
|
||||
= rewrite eq in rewrite eq2 in Calc $
|
||||
|~ indexProd (FS p) q
|
||||
~~ shift n (indexProd p q) ...( Refl )
|
||||
~~ shift n (uncurry indexProd (splitProd r)) ...( cong (shift n . uncurry indexProd) (sym eq2) )
|
||||
~~ shift n r ...( cong (shift n) (indexOfSplitProdInverse r) )
|
||||
~~ indexSum (splitSum f) ...( sym (cong indexSum eq) )
|
||||
~~ f ...( indexOfSplitSumInverse f )
|
||||
|
@ -21,11 +21,11 @@ LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:11:15--11:18, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:11:19--11:23, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:12:15--12:18, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:12:19--12:23, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG declare.data:1: Processing Term.Chk
|
||||
@ -33,14 +33,14 @@ LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:16:15--16:18, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:16:19--16:23, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:17:15--17:18, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:17:19--17:23, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:n:N}, (Term.idr:17:35--17:36, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:n:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG declare.data:1: Processing Term.Syn
|
||||
@ -48,16 +48,16 @@ LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:22:15--22:19, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:22:27--22:30, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:23:15--23:18, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:23:19--23:23, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:cut:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:24:20--24:24, Rig0))
|
||||
LOG unify.meta:5: Adding new meta ({P:vars:N}, (Term.idr:L:C--L:C, Rig0))
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
LOG unify.equal:10: Skipped unification (equal already): Type and Type
|
||||
@ -98,36 +98,36 @@ Term> Bye for now!
|
||||
LOG declare.type:1: Processing Vec.Vec
|
||||
LOG declare.def:2: Case tree for Vec.Vec: [0] ({arg:N} : (Data.Fin.Fin {arg:N}[1])) -> {arg:N}[1])
|
||||
LOG declare.type:1: Processing Vec.Nil
|
||||
LOG declare.def:2: Case tree for Vec.Nil: [0] (Prelude.Uninhabited.absurd {arg:N}[0] ?Vec.{t:N}_[{arg:N}[0]] Data.Fin.Uninhabited implementation at Data/Fin.idr:20:1--23:32)
|
||||
LOG declare.def:2: Case tree for Vec.Nil: [0] (Prelude.Uninhabited.absurd {arg:N}[0] ?Vec.{t:N}_[{arg:N}[0]] Data.Fin.Uninhabited implementation at Data/Fin.idr:L:C--L:C)
|
||||
LOG declare.type:1: Processing Vec.::
|
||||
LOG declare.def:2: Case tree for Vec.::: case {arg:N}[4] : (Data.Fin.Fin (Prelude.Types.S {arg:N}[0])) of
|
||||
{ Data.Fin.FZ {e:N} => [0] {arg:N}[3]
|
||||
| Data.Fin.FS {e:N} {e:N} => [1] ({arg:N}[5] {e:N}[1])
|
||||
}
|
||||
LOG declare.type:1: Processing Vec.test
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 2), ($resolvedN 2)] at Vec.idr:20:23--20:24
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 2), ($resolvedN 2)] at Vec.idr:L:C--L:C
|
||||
With default. Target type : Prelude.Types.Nat
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil)), (($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil)), (($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))] at Vec.idr:21:8--21:17
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil)), (($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil)), (($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))] at Vec.idr:L:C--L:C
|
||||
Target type : ({arg:N} : (Data.Fin.Fin (Prelude.Types.S (Prelude.Types.S Prelude.Types.Z)))) -> (Prelude.Types.List Prelude.Types.Nat))
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:21:9--21:11
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:L:C--L:C
|
||||
Target type : ?Vec.{a:N}_[]
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN ((:: (fromInteger 0)) Nil)) Nil), (($resolvedN ((:: (fromInteger 0)) Nil)) Nil), (($resolvedN ((:: (fromInteger 0)) Nil)) Nil)] at Vec.idr:21:8--21:17
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN ((:: (fromInteger 0)) Nil)) Nil), (($resolvedN ((:: (fromInteger 0)) Nil)) Nil), (($resolvedN ((:: (fromInteger 0)) Nil)) Nil)] at Vec.idr:L:C--L:C
|
||||
Target type : ({arg:N} : (Data.Fin.Fin ?Vec.{n:N}_[])) -> ?Vec.{a:N}_[])
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN (fromInteger 0)) Nil), (($resolvedN (fromInteger 0)) Nil), (($resolvedN (fromInteger 0)) Nil)] at Vec.idr:21:13--21:16
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [(($resolvedN (fromInteger 0)) Nil), (($resolvedN (fromInteger 0)) Nil), (($resolvedN (fromInteger 0)) Nil)] at Vec.idr:L:C--L:C
|
||||
Target type : ?Vec.{a:N}_[]
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:21:14--21:15
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:L:C--L:C
|
||||
With default. Target type : ?Vec.{a:N}_[]
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:21:13--21:16
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:L:C--L:C
|
||||
Target type : ({arg:N} : (Data.Fin.Fin ?Vec.{n:N}_[])) -> ?Vec.{a:N}_[])
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:21:14--21:15
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:L:C--L:C
|
||||
With default. Target type : ?Vec.{a:N}_[]
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:21:8--21:17
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration False [$resolvedN, $resolvedN] at Vec.idr:L:C--L:C
|
||||
Target type : ({arg:N} : (Data.Fin.Fin ?Vec.{n:N}_[])) -> ?Vec.{a:N}_[])
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration True [(($resolvedN (fromInteger 0)) Nil)] at Vec.idr:21:13--21:16
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration True [(($resolvedN (fromInteger 0)) Nil)] at Vec.idr:L:C--L:C
|
||||
Target type : (Prelude.Types.List Prelude.Types.Nat)
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:21:14--21:15
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration [($resolvedN 0), ($resolvedN 0)] at Vec.idr:L:C--L:C
|
||||
With default. Target type : Prelude.Types.Nat
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration True [$resolvedN] at Vec.idr:21:9--21:11
|
||||
LOG elab.ambiguous:5: Ambiguous elaboration True [$resolvedN] at Vec.idr:L:C--L:C
|
||||
Target type : (Prelude.Types.List Prelude.Types.Nat)
|
||||
LOG declare.def:2: Case tree for Vec.test: [0] (Vec.:: ?Vec.{n:N}_[] ?Vec.{a:N}_[] (Prelude.Types.Nil Prelude.Types.Nat) (Vec.:: Prelude.Types.Z ?Vec.{a:N}_[] (Prelude.Types.:: Prelude.Types.Nat Prelude.Types.Z (Prelude.Types.Nil Prelude.Types.Nat)) (Vec.Nil ?Vec.{a:N}_[])))
|
||||
Vec> Bye for now!
|
||||
|
@ -1,6 +1,8 @@
|
||||
echo ":q" | $1 --no-banner --no-color --console-width 0 --log unify.equal:10 --log unify:5 Term.idr \
|
||||
| sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" | sed -E "s/case in ([0-9]+)/case in N/g"
|
||||
| sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" \
|
||||
| sed -E "s/case in ([0-9]+)/case in N/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g"
|
||||
echo ":q" | $1 --no-banner --no-color --console-width 0 --log unify:3 --log elab.ambiguous:5 Vec.idr \
|
||||
| sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" | sed -E "s/case in ([0-9]+)/case in N/g"
|
||||
| sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" \
|
||||
| sed -E "s/case in ([0-9]+)/case in N/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g"
|
||||
|
||||
rm -rf build
|
@ -1,4 +1,4 @@
|
||||
import Data.Fin
|
||||
|
||||
fsprf : x === y -> FS x = FS y
|
||||
fsprf : x === y -> Fin.FS x = FS y
|
||||
fsprf p = cong _ p
|
||||
|
@ -4,7 +4,7 @@ Error: While processing right hand side of fsprf. Can't solve constraint between
|
||||
CongErr.idr:4:11--4:19
|
||||
1 | import Data.Fin
|
||||
2 |
|
||||
3 | fsprf : x === y -> FS x = FS y
|
||||
3 | fsprf : x === y -> Fin.FS x = FS y
|
||||
4 | fsprf p = cong _ p
|
||||
^^^^^^^^
|
||||
|
||||
|
@ -1,11 +1,11 @@
|
||||
1/1: Building boom (boom.idr)
|
||||
Error: While processing right hand side of with block in eraseVar. Can't solve constraint between: S (countGreater thr xs) and countGreater thr xs.
|
||||
|
||||
boom.idr:23:42--23:66
|
||||
boom.idr:23:42--23:44
|
||||
19 | eraseVar : (thr : Int) -> (ctx : Vect n Int) -> Fin n -> Maybe (Fin (countGreater thr ctx))
|
||||
20 | eraseVar thr (x :: xs) j with (x .<=. thr)
|
||||
21 | eraseVar thr (x :: xs) FZ | True = Nothing
|
||||
22 | eraseVar thr (x :: xs) FZ | False = Just FZ
|
||||
23 | eraseVar thr (x :: xs) (FS i) | True = FS <$> eraseVar thr xs i
|
||||
^^^^^^^^^^^^^^^^^^^^^^^^
|
||||
^^
|
||||
|
||||
|
Loading…
Reference in New Issue
Block a user