Idris2/tests/idris2/linear004/Stuff.idr
Edwin Brady ad632d825d Remove linearity subtyping
It's disappointing to have to do this, but I think necessary because
various issue reports have shown it to be unsound (at least as far as
inference goes) and, at the very least, confusing. This patch brings us
back to the basic rules of QTT.

On the one hand, this makes the 1 multiplicity less useful, because it
means we can't flag arguments as being used exactly once which would be
useful for optimisation purposes as well as precision in the type. On
the other hand, it removes some complexity (and a hack) from
unification, and has the advantage of being correct! Also, I still
consider the 1 multiplicity an experiment.

We can still do interesting things like protocol state tracking, which
is my primary motivation at least.

Ideally, if the 1 multiplicity is going to be more generall useful,
we'll need some kind of way of doing multiplicity polymorphism in the
future. I don't think subtyping is the way (I've pretty much always come
to regret adding some form of subtyping).

Fixes #73 (and maybe some others).
2020-12-27 19:58:35 +00:00

69 lines
1.3 KiB
Idris

-- a mini prelude
module Stuff
public export
data Bool = True | False
public export
not : Bool -> Bool
not True = False
not False = True
public export
data Maybe : Type -> Type where
Nothing : Maybe a
Just : (1 x : a) -> Maybe a
public export
data Either : Type -> Type -> Type where
Left : (1 x : a) -> Either a b
Right : (1 y : b) -> Either a b
public export
intToBool : Int -> Bool
intToBool 0 = False
intToBool x = True
public export
ifThenElse : Bool -> Lazy a -> Lazy a -> a
ifThenElse True t e = t
ifThenElse False t e = e
public export
data Nat : Type where
Z : Nat
S : (1 _ : Nat) -> Nat
public export
fromInteger : Integer -> Nat
fromInteger x = ifThenElse (intToBool (prim__eq_Integer x 0))
Z (S (fromInteger (prim__sub_Integer x 1)))
public export
plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
public export
data List : Type -> Type where
Nil : List a
(::) : (1 _ : a) -> (1 _ : List a) -> List a
public export
data Eq : a -> b -> Type where
Refl : (x : a) -> Eq x x
public export
data Unit = MkUnit
public export
data Pair : Type -> Type -> Type where
MkPair : {a, b : Type} -> a -> b -> Pair a b
public export
the : (a : Type) -> a -> a
the _ x = x