127 lines
3.5 KiB
Idris
127 lines
3.5 KiB
Idris
module Quox.Thin.List
|
|
|
|
import public Quox.Thin.Base
|
|
import public Quox.Thin.Cons
|
|
import Data.DPair
|
|
import Data.Nat
|
|
import Control.Function
|
|
|
|
%default total
|
|
|
|
||| a list of OPEs of a given outer scope size
|
|
||| (at runtime just the masks)
|
|
public export
|
|
data OPEList : Nat -> Type where
|
|
Nil : OPEList n
|
|
(::) : {mask : Nat} -> (0 ope : OPE m n mask) -> OPEList n -> OPEList n
|
|
%name OPEList opes
|
|
|
|
public export
|
|
length : OPEList n -> Nat
|
|
length [] = 0
|
|
length (_ :: opes) = S $ length opes
|
|
|
|
public export
|
|
toList : OPEList n -> List (SomeOPE n)
|
|
toList [] = []
|
|
toList (ope :: opes) = MkOPE ope :: toList opes
|
|
|
|
public export
|
|
fromList : List (SomeOPE n) -> OPEList n
|
|
fromList [] = []
|
|
fromList (MkOPE ope :: xs) = ope :: fromList xs
|
|
|
|
|
|
public export
|
|
0 Pred : Nat -> Type
|
|
Pred n = forall m, mask. OPE m n mask -> Type
|
|
|
|
public export
|
|
0 Rel : Nat -> Nat -> Type
|
|
Rel n1 n2 = forall m1, m2, mask1, mask2.
|
|
OPE m1 n1 mask1 -> OPE m2 n2 mask2 -> Type
|
|
|
|
namespace All
|
|
public export
|
|
data All : Pred n -> OPEList n -> Type where
|
|
Nil : {0 p : Pred n} -> All p []
|
|
(::) : {0 p : Pred n} -> p ope -> All p opes -> All p (ope :: opes)
|
|
%name All.All ps, qs
|
|
|
|
namespace All2
|
|
public export
|
|
data All2 : Rel n1 n2 -> OPEList n1 -> OPEList n2 -> Type where
|
|
Nil : {0 p : Rel n1 n2} -> All2 p [] []
|
|
(::) : {0 p : Rel n1 n2} -> p a b -> All2 p as bs ->
|
|
All2 p (a :: as) (b :: bs)
|
|
%name All2.All2 ps, qs
|
|
|
|
export
|
|
0 all2Length : {p : Rel m n} -> All2 p ss ts -> length ss = length ts
|
|
all2Length [] = Refl
|
|
all2Length (p :: ps) = cong S $ all2Length ps
|
|
|
|
namespace Any
|
|
public export
|
|
data Any : Pred n -> OPEList n -> Type where
|
|
Here : {0 p : Pred n} -> p ope -> Any p (ope :: opes)
|
|
There : {0 p : Pred n} -> Any p opes -> Any p (ope :: opes)
|
|
%name Any.Any p, q
|
|
|
|
export
|
|
{0 p : Pred n} -> Uninhabited (Any p []) where uninhabited _ impossible
|
|
|
|
export
|
|
all : {0 p : Pred n} ->
|
|
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> p ope) ->
|
|
(opes : OPEList n) -> All p opes
|
|
all f [] = []
|
|
all f (ope :: opes) = f ope :: all f opes
|
|
|
|
export
|
|
allDec : {0 p : Pred n} ->
|
|
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> Dec (p ope)) ->
|
|
(opes : OPEList n) -> Dec (All p opes)
|
|
allDec f [] = Yes []
|
|
allDec f (ope :: opes) = case f ope of
|
|
Yes y => case allDec f opes of
|
|
Yes ys => Yes $ y :: ys
|
|
No k => No $ \(_ :: ps) => k ps
|
|
No k => No $ \(p :: _) => k p
|
|
|
|
export
|
|
anyDec : {0 p : Pred n} ->
|
|
(forall m. {mask : Nat} -> (0 ope : OPE m n mask) -> Dec (p ope)) ->
|
|
(opes : OPEList n) -> Dec (Any p opes)
|
|
anyDec f [] = No absurd
|
|
anyDec f (ope :: opes) = case f ope of
|
|
Yes y => Yes $ Here y
|
|
No nh => case anyDec f opes of
|
|
Yes y => Yes $ There y
|
|
No nt => No $ \case Here h => nh h; There t => nt t
|
|
|
|
|
|
export
|
|
unconses : {n : Nat} -> (opes : OPEList (S n)) -> All Uncons opes
|
|
unconses = all uncons
|
|
|
|
export
|
|
heads : {n : Nat} -> (opes : OPEList (S n)) -> All (Exists . IsHead) opes
|
|
heads = all head
|
|
|
|
export
|
|
tails : {n : Nat} -> (opes : OPEList (S n)) -> All Tail opes
|
|
tails = all tail
|
|
|
|
export
|
|
tails_ : {n : Nat} -> (opes : OPEList (S n)) ->
|
|
Subset (OPEList n) (All2 IsTail opes)
|
|
tails_ [] = Element [] []
|
|
tails_ (ope :: opes) = Element _ $ (tail ope).isTail :: (tails_ opes).snd
|
|
|
|
export
|
|
conses : (heads : List Bool) -> (tails : OPEList n) ->
|
|
(0 len : length heads = length tails) =>
|
|
OPEList (S n)
|
|
conses [] [] = []
|
|
conses (h :: hs) (t :: ts) = snd (cons h t) :: conses hs ts @{inj S len}
|