2020-05-18 15:59:07 +03:00
|
|
|
module Data.Fin
|
|
|
|
|
2021-11-26 13:55:17 +03:00
|
|
|
import Control.Function
|
2021-03-01 11:29:43 +03:00
|
|
|
import Data.List1
|
Fix import loading
This was taking too long, and adding too many things, because it was
going too deep in the name of having everything accessible at the REPL
and for the compiler. So, it's done a bit differently now, only chasing
everything on a "full" load (i.e., final load at the REPL)
This has some effects:
+ As systems get bigger, load time gets better (on my machine, checking
Idris.Main now takes 52s from scratch, down from 76s)
+ You might find import errors that you didn't previously get, because
things were being imported that shouldn't have been. The new way is
correct!
An unfortunate effect is that sometimes you end up getting "undefined
name" errors even if you didn't explicitly use the name, because
sometimes a module uses a name from another module in a type, which then
gets exported, and eventually needs to be reduced. This mostly happens
because there is a compile time check that should be done which I
haven't implemented yet. That is, public export definitions should only
be allowed to use names that are also public export. I'll get to this
soon.
2020-05-27 17:49:03 +03:00
|
|
|
import public Data.Maybe
|
2021-10-07 21:21:32 +03:00
|
|
|
import public Data.Nat
|
2021-11-02 20:43:01 +03:00
|
|
|
import public Data.So
|
2020-12-04 04:49:26 +03:00
|
|
|
import Decidable.Equality.Core
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2020-06-30 01:35:34 +03:00
|
|
|
%default total
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| Numbers strictly less than some bound. The name comes from "finite sets".
|
|
|
|
|||
|
|
|
|
||| It's probably not a good idea to use `Fin` for arithmetic, and they will be
|
|
|
|
||| exceedingly inefficient at run time.
|
|
|
|
||| @ n the upper bound
|
|
|
|
public export
|
|
|
|
data Fin : (n : Nat) -> Type where
|
|
|
|
FZ : Fin (S k)
|
|
|
|
FS : Fin k -> Fin (S k)
|
|
|
|
|
2021-05-26 10:12:58 +03:00
|
|
|
||| Coerce between Fins with equal indices
|
2021-04-23 14:05:13 +03:00
|
|
|
public export
|
2021-11-25 20:07:05 +03:00
|
|
|
coerce : {n : Nat} -> (0 eq : m = n) -> Fin m -> Fin n -- TODO: make linear
|
2021-05-26 10:12:58 +03:00
|
|
|
coerce {n = S _} eq FZ = FZ
|
|
|
|
coerce {n = Z} eq FZ impossible
|
2021-11-26 13:55:17 +03:00
|
|
|
coerce {n = S _} eq (FS k) = FS (coerce (injective eq) k)
|
2021-05-26 10:12:58 +03:00
|
|
|
coerce {n = Z} eq (FS k) impossible
|
|
|
|
|
2021-11-25 20:07:05 +03:00
|
|
|
%transform "coerce-identity" coerce eq k = replace {p = Fin} eq k
|
2021-04-23 14:05:13 +03:00
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-07-05 20:52:16 +03:00
|
|
|
Uninhabited (Fin Z) where
|
2020-05-18 15:59:07 +03:00
|
|
|
uninhabited FZ impossible
|
|
|
|
uninhabited (FS f) impossible
|
|
|
|
|
2020-06-30 03:55:47 +03:00
|
|
|
export
|
|
|
|
Uninhabited (FZ = FS k) where
|
|
|
|
uninhabited Refl impossible
|
|
|
|
|
|
|
|
export
|
|
|
|
Uninhabited (FS k = FZ) where
|
|
|
|
uninhabited Refl impossible
|
|
|
|
|
2021-05-31 20:23:45 +03:00
|
|
|
export
|
|
|
|
Uninhabited (n = m) => Uninhabited (FS n = FS m) where
|
|
|
|
uninhabited Refl @{nm} = uninhabited Refl @{nm}
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-11-26 13:55:17 +03:00
|
|
|
Injective FS where
|
|
|
|
injective Refl = Refl
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
export
|
2020-07-05 20:52:16 +03:00
|
|
|
Eq (Fin n) where
|
2020-05-18 15:59:07 +03:00
|
|
|
(==) FZ FZ = True
|
|
|
|
(==) (FS k) (FS k') = k == k'
|
|
|
|
(==) _ _ = False
|
|
|
|
|
|
|
|
||| Convert a Fin to a Nat
|
|
|
|
public export
|
|
|
|
finToNat : Fin n -> Nat
|
|
|
|
finToNat FZ = Z
|
2020-07-05 20:52:16 +03:00
|
|
|
finToNat (FS k) = S $ finToNat k
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
finToNatInjective : (fm : Fin k) -> (fn : Fin k) -> (finToNat fm) = (finToNat fn) -> fm = fn
|
2020-07-05 20:52:16 +03:00
|
|
|
finToNatInjective FZ FZ _ = Refl
|
|
|
|
finToNatInjective (FS _) FZ Refl impossible
|
|
|
|
finToNatInjective FZ (FS _) Refl impossible
|
2020-05-18 15:59:07 +03:00
|
|
|
finToNatInjective (FS m) (FS n) prf =
|
2021-11-26 13:55:17 +03:00
|
|
|
cong FS $ finToNatInjective m n $ injective prf
|
|
|
|
|
|
|
|
||| `finToNat` is injective
|
|
|
|
export
|
|
|
|
Injective Fin.finToNat where
|
|
|
|
injective = (finToNatInjective _ _)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
export
|
2020-07-05 20:52:16 +03:00
|
|
|
Cast (Fin n) Nat where
|
|
|
|
cast = finToNat
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Convert a Fin to an Integer
|
|
|
|
public export
|
|
|
|
finToInteger : Fin n -> Integer
|
|
|
|
finToInteger FZ = 0
|
|
|
|
finToInteger (FS k) = 1 + finToInteger k
|
|
|
|
|
2021-08-03 16:19:17 +03:00
|
|
|
-- %builtin NaturalToInteger Data.Fin.finToInteger
|
|
|
|
|
|
|
|
export
|
|
|
|
Show (Fin n) where
|
|
|
|
show = show . finToInteger
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-07-05 20:52:16 +03:00
|
|
|
Cast (Fin n) Integer where
|
|
|
|
cast = finToInteger
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Weaken the bound on a Fin by 1
|
|
|
|
public export
|
|
|
|
weaken : Fin n -> Fin (S n)
|
|
|
|
weaken FZ = FZ
|
2020-07-05 20:52:16 +03:00
|
|
|
weaken (FS k) = FS $ weaken k
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Weaken the bound on a Fin by some amount
|
|
|
|
public export
|
2021-02-08 23:37:14 +03:00
|
|
|
weakenN : (0 n : Nat) -> Fin m -> Fin (m + n)
|
2020-05-18 15:59:07 +03:00
|
|
|
weakenN n FZ = FZ
|
2020-07-05 20:52:16 +03:00
|
|
|
weakenN n (FS f) = FS $ weakenN n f
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2020-06-12 00:14:11 +03:00
|
|
|
||| Weaken the bound on a Fin using a constructive comparison
|
|
|
|
public export
|
|
|
|
weakenLTE : Fin n -> LTE n m -> Fin m
|
|
|
|
weakenLTE FZ LTEZero impossible
|
|
|
|
weakenLTE (FS _) LTEZero impossible
|
2020-07-05 20:52:16 +03:00
|
|
|
weakenLTE FZ (LTESucc _) = FZ
|
2020-06-12 00:14:11 +03:00
|
|
|
weakenLTE (FS x) (LTESucc y) = FS $ weakenLTE x y
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| Attempt to tighten the bound on a Fin.
|
2021-06-14 12:59:49 +03:00
|
|
|
||| Return the tightened bound if there is one, else nothing.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-06-14 12:59:49 +03:00
|
|
|
strengthen : {n : _} -> Fin (S n) -> Maybe (Fin n)
|
|
|
|
strengthen {n = S _} FZ = Just FZ
|
|
|
|
strengthen {n = S _} (FS p) with (strengthen p)
|
|
|
|
strengthen (FS _) | Nothing = Nothing
|
|
|
|
strengthen (FS _) | Just q = Just $ FS q
|
|
|
|
strengthen _ = Nothing
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Add some natural number to a Fin, extending the bound accordingly
|
|
|
|
||| @ n the previous bound
|
|
|
|
||| @ m the number to increase the Fin by
|
|
|
|
public export
|
|
|
|
shift : (m : Nat) -> Fin n -> Fin (m + n)
|
|
|
|
shift Z f = f
|
2020-07-05 20:52:16 +03:00
|
|
|
shift (S m) f = FS $ shift m f
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| The largest element of some Fin type
|
|
|
|
public export
|
|
|
|
last : {n : _} -> Fin (S n)
|
|
|
|
last {n=Z} = FZ
|
|
|
|
last {n=S _} = FS last
|
|
|
|
|
2021-03-01 11:29:43 +03:00
|
|
|
||| All of the Fin elements
|
|
|
|
public export
|
|
|
|
allFins : (n : Nat) -> List1 (Fin (S n))
|
|
|
|
allFins Z = FZ ::: []
|
|
|
|
allFins (S n) = FZ ::: map FS (forget (allFins n))
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-07-05 20:52:16 +03:00
|
|
|
Ord (Fin n) where
|
2020-05-18 15:59:07 +03:00
|
|
|
compare FZ FZ = EQ
|
|
|
|
compare FZ (FS _) = LT
|
|
|
|
compare (FS _) FZ = GT
|
|
|
|
compare (FS x) (FS y) = compare x y
|
|
|
|
|
2021-10-07 21:21:32 +03:00
|
|
|
public export
|
|
|
|
natToFinLT : (x : Nat) -> {0 n : Nat} ->
|
|
|
|
{auto 0 prf : x `LT` n} ->
|
|
|
|
Fin n
|
|
|
|
natToFinLT Z {prf = LTESucc _} = FZ
|
|
|
|
natToFinLT (S k) {prf = LTESucc _} = FS $ natToFinLT k
|
|
|
|
|
2021-11-02 20:43:01 +03:00
|
|
|
public export
|
|
|
|
natToFinLt : (x : Nat) -> {n : Nat} ->
|
|
|
|
{auto 0 prf : So (x < n)} ->
|
|
|
|
Fin n
|
|
|
|
natToFinLt Z {n = S _} = FZ
|
|
|
|
natToFinLt (S k) {n = S _} = FS $ natToFinLt k
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
public export
|
|
|
|
natToFin : Nat -> (n : Nat) -> Maybe (Fin n)
|
2021-10-07 21:21:32 +03:00
|
|
|
natToFin x n = case isLT x n of
|
|
|
|
Yes prf => Just $ natToFinLT x
|
|
|
|
No contra => Nothing
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Convert an `Integer` to a `Fin`, provided the integer is within bounds.
|
|
|
|
||| @n The upper bound of the Fin
|
|
|
|
public export
|
|
|
|
integerToFin : Integer -> (n : Nat) -> Maybe (Fin n)
|
|
|
|
integerToFin x Z = Nothing -- make sure 'n' is concrete, to save reduction!
|
|
|
|
integerToFin x n = if x >= 0 then natToFin (fromInteger x) n else Nothing
|
|
|
|
|
2021-10-07 21:21:32 +03:00
|
|
|
public export
|
|
|
|
maybeLTE : (x : Nat) -> (y : Nat) -> Maybe (x `LTE` y)
|
|
|
|
maybeLTE Z _ = Just LTEZero
|
|
|
|
maybeLTE (S x) (S y) = LTESucc <$> maybeLTE x y
|
|
|
|
maybeLTE _ _ = Nothing
|
|
|
|
|
|
|
|
public export
|
|
|
|
maybeLT : (x : Nat) -> (y : Nat) -> Maybe (x `LT` y)
|
|
|
|
maybeLT x y = maybeLTE (S x) y
|
|
|
|
|
2021-12-13 16:47:53 +03:00
|
|
|
public export
|
|
|
|
finFromInteger : (x : Integer) -> {n : Nat} ->
|
|
|
|
{auto 0 prf : So (fromInteger x < n)} ->
|
|
|
|
Fin n
|
|
|
|
finFromInteger x = natToFinLt (fromInteger x)
|
|
|
|
|
|
|
|
-- Direct comparison between `Integer` and `Nat`.
|
|
|
|
-- We cannot convert the `Nat` to `Integer`, because that will not work with open terms;
|
|
|
|
-- but we cannot convert the `Integer` to `Nat` either, because that slows down typechecking
|
|
|
|
-- even when the function is unused (issue #2032)
|
|
|
|
public export
|
|
|
|
integerLessThanNat : Integer -> Nat -> Bool
|
|
|
|
integerLessThanNat x n with (x < the Integer 0)
|
|
|
|
integerLessThanNat _ _ | True = True -- if `x < 0` then `x < n` for any `n : Nat`
|
|
|
|
integerLessThanNat x (S m) | False = integerLessThanNat (x-1) m -- recursive case
|
|
|
|
integerLessThanNat x Z | False = False -- `x >= 0` contradicts `x < Z`
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| Allow overloading of Integer literals for Fin.
|
|
|
|
||| @ x the Integer that the user typed
|
|
|
|
||| @ prf an automatically-constructed proof that `x` is in bounds
|
|
|
|
public export
|
|
|
|
fromInteger : (x : Integer) -> {n : Nat} ->
|
2021-12-13 16:47:53 +03:00
|
|
|
{auto 0 prf : So (integerLessThanNat x n)} ->
|
2020-05-18 15:59:07 +03:00
|
|
|
Fin n
|
2021-12-13 16:47:53 +03:00
|
|
|
fromInteger x = finFromInteger x {prf = lemma prf} where
|
|
|
|
-- to be minimally invasive, we just call the previous implementation.
|
|
|
|
-- however, having a different proof obligation resolves #2032
|
|
|
|
0 lemma : {x : Integer} -> {n : Nat} -> So (integerLessThanNat x n) -> So (fromInteger {ty=Nat} x < n)
|
|
|
|
lemma oh = believe_me oh
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-08-03 16:19:17 +03:00
|
|
|
-- %builtin IntegerToNatural Data.Fin.fromInteger
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
||| Convert an Integer to a Fin in the required bounds/
|
|
|
|
||| This is essentially a composition of `mod` and `fromInteger`
|
|
|
|
public export
|
|
|
|
restrict : (n : Nat) -> Integer -> Fin (S n)
|
|
|
|
restrict n val = let val' = assert_total (abs (mod val (cast (S n)))) in
|
|
|
|
-- reasoning about primitives, so we need the
|
|
|
|
-- 'believe_me'. It's fine because val' must be
|
|
|
|
-- in the right range
|
|
|
|
fromInteger {n = S n} val'
|
|
|
|
{prf = believe_me {a=IsJust (Just val')} ItIsJust}
|
|
|
|
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- DecEq
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
|
2021-03-18 14:17:48 +03:00
|
|
|
public export
|
2020-07-05 20:52:16 +03:00
|
|
|
DecEq (Fin n) where
|
2020-05-18 15:59:07 +03:00
|
|
|
decEq FZ FZ = Yes Refl
|
2020-07-05 20:52:16 +03:00
|
|
|
decEq FZ (FS f) = No absurd
|
|
|
|
decEq (FS f) FZ = No absurd
|
2020-05-18 15:59:07 +03:00
|
|
|
decEq (FS f) (FS f')
|
|
|
|
= case decEq f f' of
|
|
|
|
Yes p => Yes $ cong FS p
|
2021-11-26 13:55:17 +03:00
|
|
|
No p => No $ p . injective
|
2021-04-23 14:05:13 +03:00
|
|
|
|
|
|
|
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)
|
|
|
|
|
2021-05-26 10:12:58 +03:00
|
|
|
||| Pointwise equality is compatible with coerce
|
2021-04-23 14:05:13 +03:00
|
|
|
export
|
2021-05-26 10:12:58 +03:00
|
|
|
coerceEq : {k : Fin m} -> (0 eq : m = n) -> coerce eq k ~~~ k
|
|
|
|
coerceEq {k = FZ} Refl = FZ
|
|
|
|
coerceEq {k = FS k} Refl = FS (coerceEq _)
|
2021-04-23 14:05:13 +03:00
|
|
|
|
2021-05-26 10:12:58 +03:00
|
|
|
||| The actual proof used by coerce is irrelevant
|
2021-04-23 14:05:13 +03:00
|
|
|
export
|
2021-05-26 10:12:58 +03:00
|
|
|
congCoerce : {0 n, q : Nat} -> {k : Fin m} -> {l : Fin p} ->
|
|
|
|
{0 eq1 : m = n} -> {0 eq2 : p = q} ->
|
|
|
|
k ~~~ l ->
|
|
|
|
coerce eq1 k ~~~ coerce eq2 l
|
|
|
|
congCoerce eq
|
|
|
|
= transitive (coerceEq _)
|
|
|
|
$ transitive eq $ symmetric $ coerceEq _
|
2021-04-23 14:05:13 +03:00
|
|
|
|
|
|
|
||| Last is congruent wrt index equality
|
|
|
|
export
|
2021-05-07 16:55:15 +03:00
|
|
|
congLast : {m : Nat} -> (0 _ : m = n) -> last {n=m} ~~~ last {n}
|
|
|
|
congLast Refl = reflexive
|
2021-04-23 14:05:13 +03:00
|
|
|
|
|
|
|
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} ->
|
2021-05-07 16:55:15 +03:00
|
|
|
(0 eq : m === n) -> k ~~~ l -> k === rewrite eq in l
|
2021-04-23 14:05:13 +03:00
|
|
|
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)
|