mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-26 05:01:34 +03:00
Add linear pair/dependent pair to the prelude
I'm playing with some linear structures and finding these useful a lot, so good to have a consistent syntax for it. '#' is chosen because it's short, looks a bit like a cross if you look at it from the right angle (!) and so as not to clash with '@@' in preorder reasoning syntax.
This commit is contained in:
parent
361d2e4d88
commit
c9b20911e1
@ -83,7 +83,7 @@ access the resource directly:
|
|||||||
.. code-block:: idris
|
.. code-block:: idris
|
||||||
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
data Res : (a : Type) -> (a -> Type) -> Type where
|
||||||
(@@) : (val : a) -> (1 resource : r val) -> Res a r
|
(#) : (val : a) -> (1 resource : r val) -> Res a r
|
||||||
|
|
||||||
login : (1 s : Store LoggedOut) -> (password : String) ->
|
login : (1 s : Store LoggedOut) -> (password : String) ->
|
||||||
Res Bool (\ok => Store (if ok then LoggedIn else LoggedOut))
|
Res Bool (\ok => Store (if ok then LoggedIn else LoggedOut))
|
||||||
@ -91,7 +91,7 @@ access the resource directly:
|
|||||||
readSecret : (1 s : Store LoggedIn) ->
|
readSecret : (1 s : Store LoggedIn) ->
|
||||||
Res String (const (Store LoggedIn))
|
Res String (const (Store LoggedIn))
|
||||||
|
|
||||||
``Res`` is defined in ``Control.App`` since it is commonly useful. It is a
|
``Res`` is defined in the Prelude, since it is commonly useful. It is a
|
||||||
dependent pair type, which associates a value with a linear resource.
|
dependent pair type, which associates a value with a linear resource.
|
||||||
We'll leave the other definitions abstract, for the purposes of this
|
We'll leave the other definitions abstract, for the purposes of this
|
||||||
introductory example.
|
introductory example.
|
||||||
@ -108,10 +108,10 @@ secret data. It uses ``let (>>=) = bindL`` to redefine
|
|||||||
do putStr "Password: "
|
do putStr "Password: "
|
||||||
password <- getStr
|
password <- getStr
|
||||||
connect $ \s =>
|
connect $ \s =>
|
||||||
do let True @@ s = login s password
|
do let True # s = login s password
|
||||||
| False @@ s => do putStrLn "Wrong password"
|
| False # s => do putStrLn "Wrong password"
|
||||||
disconnect s
|
disconnect s
|
||||||
let str @@ s = readSecret s
|
let str # s = readSecret s
|
||||||
putStrLn $ "Secret: " ++ show str
|
putStrLn $ "Secret: " ++ show str
|
||||||
let s = logout s
|
let s = logout s
|
||||||
disconnect s
|
disconnect s
|
||||||
@ -237,10 +237,10 @@ hard coded password and internal data:
|
|||||||
|
|
||||||
login (MkStore str) pwd
|
login (MkStore str) pwd
|
||||||
= if pwd == "Mornington Crescent"
|
= if pwd == "Mornington Crescent"
|
||||||
then pure1 (True @@ MkStore str)
|
then pure1 (True # MkStore str)
|
||||||
else pure1 (False @@ MkStore str)
|
else pure1 (False # MkStore str)
|
||||||
logout (MkStore str) = pure1 (MkStore str)
|
logout (MkStore str) = pure1 (MkStore str)
|
||||||
readSecret (MkStore str) = pure1 (str @@ MkStore str)
|
readSecret (MkStore str) = pure1 (str # MkStore str)
|
||||||
|
|
||||||
disconnect (MkStore _)
|
disconnect (MkStore _)
|
||||||
= putStrLn "Door destroyed"
|
= putStrLn "Door destroyed"
|
||||||
|
@ -339,8 +339,6 @@ HasErr Void e => PrimIO e where
|
|||||||
$ \_ =>
|
$ \_ =>
|
||||||
MkAppRes (Right ())
|
MkAppRes (Right ())
|
||||||
|
|
||||||
infix 5 @@
|
|
||||||
|
|
||||||
export
|
export
|
||||||
new1 : t -> (1 p : State tag t e => App1 {u} e a) -> App1 {u} e a
|
new1 : t -> (1 p : State tag t e => App1 {u} e a) -> App1 {u} e a
|
||||||
new1 val prog
|
new1 val prog
|
||||||
@ -349,10 +347,6 @@ new1 val prog
|
|||||||
MkApp1 res = prog @{st} in
|
MkApp1 res = prog @{st} in
|
||||||
res
|
res
|
||||||
|
|
||||||
public export
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (val : a) -> (1 r : t val) -> Res a t
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data FileEx = GenericFileEx Int -- errno
|
data FileEx = GenericFileEx Int -- errno
|
||||||
| FileReadError
|
| FileReadError
|
||||||
|
@ -67,22 +67,36 @@ snd (x, y) = y
|
|||||||
-- This directive tells auto implicit search what to use to look inside pairs
|
-- This directive tells auto implicit search what to use to look inside pairs
|
||||||
%pair Pair fst snd
|
%pair Pair fst snd
|
||||||
|
|
||||||
||| Dependent pairs aid in the construction of dependent types by providing
|
infix 5 #
|
||||||
||| evidence that some value resides in the type.
|
|
||||||
|||
|
||| A pair type for use in operations on linear resources, which return a
|
||||||
||| Formally, speaking, dependent pairs represent existential quantification -
|
||| value and an updated resource
|
||||||
||| they consist of a witness for the existential claim and a proof that the
|
public export
|
||||||
||| property holds for it.
|
data LPair : Type -> Type -> Type where
|
||||||
|||
|
(#) : (x : a) -> (1 _ : b) -> LPair a b
|
||||||
||| @ a the value to place in the type.
|
|
||||||
||| @ p the dependent type that requires the value.
|
|
||||||
namespace DPair
|
namespace DPair
|
||||||
|
||| Dependent pairs aid in the construction of dependent types by providing
|
||||||
|
||| evidence that some value resides in the type.
|
||||||
|
|||
|
||||||
|
||| Formally, speaking, dependent pairs represent existential quantification -
|
||||||
|
||| they consist of a witness for the existential claim and a proof that the
|
||||||
|
||| property holds for it.
|
||||||
|
|||
|
||||||
|
||| @ a the value to place in the type.
|
||||||
|
||| @ p the dependent type that requires the value.
|
||||||
public export
|
public export
|
||||||
record DPair a (p : a -> Type) where
|
record DPair a (p : a -> Type) where
|
||||||
constructor MkDPair
|
constructor MkDPair
|
||||||
fst : a
|
fst : a
|
||||||
snd : p fst
|
snd : p fst
|
||||||
|
|
||||||
|
||| A dependent variant of LPair, pairing a result value with a resource
|
||||||
|
||| that depends on the result value
|
||||||
|
public export
|
||||||
|
data Res : (a : Type) -> (a -> Type) -> Type where
|
||||||
|
(#) : (val : a) -> (1 r : t val) -> Res a t
|
||||||
|
|
||||||
-- The empty type
|
-- The empty type
|
||||||
|
|
||||||
||| The empty type, also known as the trivially false proposition.
|
||| The empty type, also known as the trivially false proposition.
|
||||||
|
@ -2,13 +2,13 @@
|
|||||||
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
||||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1)
|
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1)
|
||||||
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 8 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 8 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
||||||
0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 38)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:260} a)")))))) 1)
|
0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 38)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:269} a)")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:261}_[]")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:270}_[]")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:260}_[] ?{_:261}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:269}_[] ?{_:270}_[])")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 21)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:262}_[] ?{_:261}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 21)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:271}_[] ?{_:270}_[])")))))) 1)
|
||||||
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 7 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 7 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 14)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:251}_[] ?{_:250}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 14)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:260}_[] ?{_:259}_[])")))))) 1)
|
||||||
0001ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 6 1)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Nat} -> {0 a : Type} -> {0 n : Prelude.Nat} -> (({arg:241} : (Main.Vect n[0] a[1])) -> (({arg:242} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.+ Prelude.Nat Prelude.Num implementation at Prelude.idr:849:1--856:1 n[2] m[4]) a[3])))")))))) 1)
|
0001ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 6 1)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Nat} -> {0 a : Type} -> {0 n : Prelude.Nat} -> (({arg:250} : (Main.Vect n[0] a[1])) -> (({arg:251} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.+ Prelude.Nat Prelude.Num implementation at Prelude.idr:849:1--856:1 n[2] m[4]) a[3])))")))))) 1)
|
||||||
0000cb(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 6 1)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1)
|
0000cb(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 6 1)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 42)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 42)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
||||||
|
@ -2,13 +2,13 @@
|
|||||||
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
||||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1)
|
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1)
|
||||||
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 8 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 8 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
||||||
0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 38)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:260} a)")))))) 1)
|
0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 38)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:269} a)")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:261}_[]")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:270}_[]")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:260}_[] ?{_:261}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:269}_[] ?{_:270}_[])")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 21)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:262}_[] ?{_:261}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 21)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:271}_[] ?{_:270}_[])")))))) 1)
|
||||||
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 7 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
0000d8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 7 1)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1)
|
||||||
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 14)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:251}_[] ?{_:250}_[])")))))) 1)
|
0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 14)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:260}_[] ?{_:259}_[])")))))) 1)
|
||||||
0001ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 6 1)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Nat} -> {0 a : Type} -> {0 n : Prelude.Nat} -> (({arg:241} : (Main.Vect n[0] a[1])) -> (({arg:242} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.+ Prelude.Nat Prelude.Num implementation at Prelude.idr:849:1--856:1 n[2] m[4]) a[3])))")))))) 1)
|
0001ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 6 1)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Nat} -> {0 a : Type} -> {0 n : Prelude.Nat} -> (({arg:250} : (Main.Vect n[0] a[1])) -> (({arg:251} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.+ Prelude.Nat Prelude.Num implementation at Prelude.idr:849:1--856:1 n[2] m[4]) a[3])))")))))) 1)
|
||||||
0000cb(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 6 1)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1)
|
0000cb(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 6 1)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
||||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 42)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 42)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Nat")))))) 1)
|
||||||
|
@ -1,6 +1,6 @@
|
|||||||
1/1: Building wcov (wcov.idr)
|
1/1: Building wcov (wcov.idr)
|
||||||
Main> Main.tfoo is total
|
Main> Main.tfoo is total
|
||||||
Main> Main.wfoo is total
|
Main> Main.wfoo is total
|
||||||
Main> Main.wbar is not covering due to call to function Main.with block in 1372
|
Main> Main.wbar is not covering due to call to function Main.with block in 1376
|
||||||
Main> Main.wbar1 is total
|
Main> Main.wbar1 is total
|
||||||
Main> Bye for now!
|
Main> Bye for now!
|
||||||
|
@ -1,3 +1,3 @@
|
|||||||
1/1: Building casetot (casetot.idr)
|
1/1: Building casetot (casetot.idr)
|
||||||
casetot.idr:12:1--13:1:main is not covering:
|
casetot.idr:12:1--13:1:main is not covering:
|
||||||
Calls non covering function Main.case block in 2071(287)
|
Calls non covering function Main.case block in 2075(296)
|
||||||
|
@ -22,11 +22,6 @@ One = Use Once
|
|||||||
Any : (Type -> Type) -> Type -> Type
|
Any : (Type -> Type) -> Type -> Type
|
||||||
Any = Use Many
|
Any = Use Many
|
||||||
|
|
||||||
infix 2 @@
|
|
||||||
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (x : a) -> (1 res : r x) -> Res a r
|
|
||||||
|
|
||||||
data DoorState = Closed | Open
|
data DoorState = Closed | Open
|
||||||
|
|
||||||
data Door : DoorState -> Type where
|
data Door : DoorState -> Type where
|
||||||
@ -57,7 +52,7 @@ doorProg2
|
|||||||
r <- openDoor d
|
r <- openDoor d
|
||||||
let x = 42
|
let x = 42
|
||||||
case r of
|
case r of
|
||||||
(res @@ d) => ?now_1
|
(res # d) => ?now_1
|
||||||
|
|
||||||
doorProg3 : Any m ()
|
doorProg3 : Any m ()
|
||||||
doorProg3
|
doorProg3
|
||||||
@ -65,5 +60,5 @@ doorProg3
|
|||||||
r <- openDoor d
|
r <- openDoor d
|
||||||
let x = 42
|
let x = 42
|
||||||
case r of
|
case r of
|
||||||
(True @@ d) => ?now_2
|
(True # d) => ?now_2
|
||||||
(False @@ d) => ?now_3
|
(False # d) => ?now_3
|
||||||
|
@ -1,7 +1,7 @@
|
|||||||
1/1: Building Door (Door.idr)
|
1/1: Building Door (Door.idr)
|
||||||
Main> (y @@ res) => ?now_4
|
Main> (val # y) => ?now_4
|
||||||
Main> (True @@ d) => ?now_4
|
Main> (True # d) => ?now_4
|
||||||
(False @@ d) => ?now_5
|
(False # d) => ?now_5
|
||||||
Main> 0 m : Type -> Type
|
Main> 0 m : Type -> Type
|
||||||
1 d : Door Open
|
1 d : Door Open
|
||||||
x : Integer
|
x : Integer
|
||||||
|
@ -1,5 +1,5 @@
|
|||||||
:cs 52 15 what
|
:cs 47 15 what
|
||||||
:cs 60 16 res
|
:cs 55 16 res
|
||||||
:t now_2
|
:t now_2
|
||||||
:t now_3
|
:t now_3
|
||||||
:q
|
:q
|
||||||
|
@ -30,9 +30,3 @@ One = Use Once
|
|||||||
public export
|
public export
|
||||||
Any : (Type -> Type) -> Type -> Type
|
Any : (Type -> Type) -> Type -> Type
|
||||||
Any = Use Many
|
Any = Use Many
|
||||||
|
|
||||||
infix 2 @@
|
|
||||||
|
|
||||||
public export
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (x : a) -> (1 res : r x) -> Res a r
|
|
||||||
|
@ -56,12 +56,6 @@ public export
|
|||||||
Any : (Type -> Type) -> Type -> Type
|
Any : (Type -> Type) -> Type -> Type
|
||||||
Any m = Lin m Many
|
Any m = Lin m Many
|
||||||
|
|
||||||
infix 2 @@
|
|
||||||
|
|
||||||
public export
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (val : a) -> (1 resource : r val) -> Res a r
|
|
||||||
|
|
||||||
data DoorState = Closed | Open
|
data DoorState = Closed | Open
|
||||||
|
|
||||||
-- changes start here
|
-- changes start here
|
||||||
|
@ -22,11 +22,6 @@
|
|||||||
> Any : (Type -> Type) -> Type -> Type
|
> Any : (Type -> Type) -> Type -> Type
|
||||||
> Any = Use Many
|
> Any = Use Many
|
||||||
|
|
||||||
> infix 2 @@
|
|
||||||
|
|
||||||
> data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
> (@@) : (x : a) -> (1 res : r x) -> Res a r
|
|
||||||
|
|
||||||
> data DoorState = Closed | Open
|
> data DoorState = Closed | Open
|
||||||
|
|
||||||
> data Door : DoorState -> Type where
|
> data Door : DoorState -> Type where
|
||||||
@ -57,7 +52,7 @@
|
|||||||
> r <- openDoor d
|
> r <- openDoor d
|
||||||
> let x = 42
|
> let x = 42
|
||||||
> case r of
|
> case r of
|
||||||
> (res @@ d) => ?now_1
|
> (res # d) => ?now_1
|
||||||
|
|
||||||
> doorProg3 : Any m ()
|
> doorProg3 : Any m ()
|
||||||
> doorProg3
|
> doorProg3
|
||||||
@ -65,5 +60,5 @@
|
|||||||
> r <- openDoor d
|
> r <- openDoor d
|
||||||
> let x = 42
|
> let x = 42
|
||||||
> case r of
|
> case r of
|
||||||
> (True @@ d) => ?now_2
|
> (True # d) => ?now_2
|
||||||
> (False @@ d) => ?now_3
|
> (False # d) => ?now_3
|
||||||
|
@ -1,7 +1,7 @@
|
|||||||
1/1: Building Door (Door.lidr)
|
1/1: Building Door (Door.lidr)
|
||||||
Main> > (y @@ res) => ?now_4
|
Main> > (val # y) => ?now_4
|
||||||
Main> > (True @@ d) => ?now_4
|
Main> > (True # d) => ?now_4
|
||||||
> (False @@ d) => ?now_5
|
> (False # d) => ?now_5
|
||||||
Main> 0 m : Type -> Type
|
Main> 0 m : Type -> Type
|
||||||
1 d : Door Open
|
1 d : Door Open
|
||||||
x : Integer
|
x : Integer
|
||||||
|
@ -1,5 +1,5 @@
|
|||||||
:cs 52 17 what
|
:cs 47 17 what
|
||||||
:cs 60 18 res
|
:cs 55 18 res
|
||||||
:t now_2
|
:t now_2
|
||||||
:t now_3
|
:t now_3
|
||||||
:q
|
:q
|
||||||
|
@ -134,11 +134,11 @@ tryRecv (MkChannel lock cond_lock cond local remote)
|
|||||||
case dequeue rq of
|
case dequeue rq of
|
||||||
Nothing =>
|
Nothing =>
|
||||||
do lift $ mutexRelease lock
|
do lift $ mutexRelease lock
|
||||||
pure (Nothing @@ MkChannel lock cond_lock cond local remote)
|
pure (Nothing # MkChannel lock cond_lock cond local remote)
|
||||||
Just (rq', Entry {any} val) =>
|
Just (rq', Entry {any} val) =>
|
||||||
do lift $ writeIORef local rq'
|
do lift $ writeIORef local rq'
|
||||||
lift $ mutexRelease lock
|
lift $ mutexRelease lock
|
||||||
pure (Just (believe_me {a=any} val) @@
|
pure (Just (believe_me {a=any} val) #
|
||||||
MkChannel lock cond_lock cond local remote)
|
MkChannel lock cond_lock cond local remote)
|
||||||
|
|
||||||
-- blocks until the message is there
|
-- blocks until the message is there
|
||||||
@ -158,7 +158,7 @@ recv (MkChannel lock cond_lock cond local remote)
|
|||||||
Just (rq', Entry {any} val) =>
|
Just (rq', Entry {any} val) =>
|
||||||
do lift $ writeIORef local rq'
|
do lift $ writeIORef local rq'
|
||||||
lift $ mutexRelease lock
|
lift $ mutexRelease lock
|
||||||
pure (believe_me {a=any} val @@
|
pure (believe_me {a=any} val #
|
||||||
MkChannel lock cond_lock cond local remote)
|
MkChannel lock cond_lock cond local remote)
|
||||||
|
|
||||||
export
|
export
|
||||||
@ -180,14 +180,14 @@ timeoutRecv (MkChannel lock cond_lock cond local remote) timeout
|
|||||||
lift $ mutexAcquire cond_lock
|
lift $ mutexAcquire cond_lock
|
||||||
lift $ conditionWaitTimeout cond cond_lock timeout
|
lift $ conditionWaitTimeout cond cond_lock timeout
|
||||||
lift $ mutexRelease cond_lock
|
lift $ mutexRelease cond_lock
|
||||||
res @@ chan <- tryRecv {ty} {next} (MkChannel lock cond_lock cond local remote)
|
res # chan <- tryRecv {ty} {next} (MkChannel lock cond_lock cond local remote)
|
||||||
case res of
|
case res of
|
||||||
Nothing => pure (Nothing @@ chan)
|
Nothing => pure (Nothing # chan)
|
||||||
Just res => pure (Just res @@ chan)
|
Just res => pure (Just res # chan)
|
||||||
Just (rq', Entry {any} val) =>
|
Just (rq', Entry {any} val) =>
|
||||||
do lift $ writeIORef local rq'
|
do lift $ writeIORef local rq'
|
||||||
lift $ mutexRelease lock
|
lift $ mutexRelease lock
|
||||||
pure (Just (believe_me {a=any} val) @@
|
pure (Just (believe_me {a=any} val) #
|
||||||
MkChannel lock cond_lock cond local remote)
|
MkChannel lock cond_lock cond local remote)
|
||||||
|
|
||||||
export
|
export
|
||||||
|
@ -54,13 +54,6 @@ public export
|
|||||||
Any : (Type -> Type) -> Type -> Type
|
Any : (Type -> Type) -> Type -> Type
|
||||||
Any m = Lin m Many
|
Any m = Lin m Many
|
||||||
|
|
||||||
infix 2 @@
|
|
||||||
|
|
||||||
public export
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (val : a) -> (1 resource : r val) -> Res a r
|
|
||||||
|
|
||||||
|
|
||||||
{-
|
{-
|
||||||
data DoorState = Closed | Open
|
data DoorState = Closed | Open
|
||||||
|
|
||||||
|
@ -18,12 +18,12 @@ Utils
|
|||||||
|
|
||||||
utilServer : (1 chan : Server Utils) -> Any IO ()
|
utilServer : (1 chan : Server Utils) -> Any IO ()
|
||||||
utilServer chan
|
utilServer chan
|
||||||
= do cmd @@ chan <- recv chan
|
= do cmd # chan <- recv chan
|
||||||
case cmd of
|
case cmd of
|
||||||
Add => do (x, y) @@ chan <- recv chan
|
Add => do (x, y) # chan <- recv chan
|
||||||
chan <- send chan (x + y)
|
chan <- send chan (x + y)
|
||||||
close chan
|
close chan
|
||||||
Append => do (x, y) @@ chan <- recv chan
|
Append => do (x, y) # chan <- recv chan
|
||||||
chan <- send chan (x ++ y)
|
chan <- send chan (x ++ y)
|
||||||
close chan
|
close chan
|
||||||
|
|
||||||
@ -35,7 +35,7 @@ MakeUtils = do cmd <- Request Bool
|
|||||||
|
|
||||||
sendUtils : (1 chan : Server MakeUtils) -> Any IO ()
|
sendUtils : (1 chan : Server MakeUtils) -> Any IO ()
|
||||||
sendUtils chan
|
sendUtils chan
|
||||||
= do cmd @@ chan <- recv chan
|
= do cmd # chan <- recv chan
|
||||||
if cmd
|
if cmd
|
||||||
then do cchan <- Channel.fork utilServer
|
then do cchan <- Channel.fork utilServer
|
||||||
chan <- send chan cchan
|
chan <- send chan cchan
|
||||||
@ -46,7 +46,7 @@ getUtilsChan : (1 chan : Client MakeUtils) ->
|
|||||||
One IO (Client Utils, Client MakeUtils)
|
One IO (Client Utils, Client MakeUtils)
|
||||||
getUtilsChan chan
|
getUtilsChan chan
|
||||||
= do chan <- send chan True
|
= do chan <- send chan True
|
||||||
cchan @@ chan <- recv chan
|
cchan # chan <- recv chan
|
||||||
pure (cchan, chan)
|
pure (cchan, chan)
|
||||||
|
|
||||||
closeUtilsChan : (1 chan : Client MakeUtils) ->
|
closeUtilsChan : (1 chan : Client MakeUtils) ->
|
||||||
@ -69,12 +69,12 @@ doThings
|
|||||||
uchan1 <- send uchan1 Add
|
uchan1 <- send uchan1 Add
|
||||||
uchan2 <- send uchan2 Append
|
uchan2 <- send uchan2 Append
|
||||||
uchan2 <- send uchan2 ("aaa", "bbb")
|
uchan2 <- send uchan2 ("aaa", "bbb")
|
||||||
res @@ uchan2 <- recv uchan2
|
res # uchan2 <- recv uchan2
|
||||||
close uchan2
|
close uchan2
|
||||||
lift $ printLn res
|
lift $ printLn res
|
||||||
|
|
||||||
uchan1 <- send uchan1 (40, 54)
|
uchan1 <- send uchan1 (40, 54)
|
||||||
res @@ uchan1 <- recv uchan1
|
res # uchan1 <- recv uchan1
|
||||||
close uchan1
|
close uchan1
|
||||||
|
|
||||||
lift $ printLn res
|
lift $ printLn res
|
||||||
|
@ -18,14 +18,14 @@ testClient chan
|
|||||||
lift $ putStrLn "Sending value"
|
lift $ putStrLn "Sending value"
|
||||||
chan <- send chan False
|
chan <- send chan False
|
||||||
lift $ putStrLn "Sent"
|
lift $ putStrLn "Sent"
|
||||||
c @@ chan <- recv chan
|
c # chan <- recv chan
|
||||||
lift $ putStrLn ("Result: " ++ c)
|
lift $ putStrLn ("Result: " ++ c)
|
||||||
close chan
|
close chan
|
||||||
|
|
||||||
testServer : (1 chan : Server TestProto) -> Any IO ()
|
testServer : (1 chan : Server TestProto) -> Any IO ()
|
||||||
testServer chan
|
testServer chan
|
||||||
= do lift $ putStrLn "Waiting"
|
= do lift $ putStrLn "Waiting"
|
||||||
cmd @@ chan <- recv chan
|
cmd # chan <- recv chan
|
||||||
lift $ putStrLn ("Received " ++ show cmd)
|
lift $ putStrLn ("Received " ++ show cmd)
|
||||||
lift $ sleep 1
|
lift $ sleep 1
|
||||||
lift $ putStrLn "Sending answer"
|
lift $ putStrLn "Sending answer"
|
||||||
|
@ -339,12 +339,6 @@ HasErr Void e => PrimIO e where
|
|||||||
$ \_ =>
|
$ \_ =>
|
||||||
MkAppRes (Right ())
|
MkAppRes (Right ())
|
||||||
|
|
||||||
infix 5 @@
|
|
||||||
|
|
||||||
public export
|
|
||||||
data Res : (a : Type) -> (a -> Type) -> Type where
|
|
||||||
(@@) : (val : a) -> (1 r : t val) -> Res a t
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data FileEx = GenericFileEx Int -- errno
|
data FileEx = GenericFileEx Int -- errno
|
||||||
| FileReadError
|
| FileReadError
|
||||||
|
@ -23,10 +23,10 @@ Has [Console] e => StoreI e where
|
|||||||
|
|
||||||
login (MkStore str) pwd
|
login (MkStore str) pwd
|
||||||
= if pwd == "Mornington Crescent"
|
= if pwd == "Mornington Crescent"
|
||||||
then pure1 (True @@ MkStore str)
|
then pure1 (True # MkStore str)
|
||||||
else pure1 (False @@ MkStore str)
|
else pure1 (False # MkStore str)
|
||||||
logout (MkStore str) = pure1 (MkStore str)
|
logout (MkStore str) = pure1 (MkStore str)
|
||||||
readSecret (MkStore str) = pure1 (str @@ MkStore str)
|
readSecret (MkStore str) = pure1 (str # MkStore str)
|
||||||
|
|
||||||
disconnect (MkStore _)
|
disconnect (MkStore _)
|
||||||
= putStrLn "Door destroyed"
|
= putStrLn "Door destroyed"
|
||||||
@ -38,11 +38,11 @@ storeProg
|
|||||||
s <- connect
|
s <- connect
|
||||||
app $ putStr "Password: "
|
app $ putStr "Password: "
|
||||||
pwd <- app $ getStr
|
pwd <- app $ getStr
|
||||||
True @@ s <- login s pwd
|
True # s <- login s pwd
|
||||||
| False @@ s => do app $ putStrLn "Login failed"
|
| False # s => do app $ putStrLn "Login failed"
|
||||||
app $ disconnect s
|
app $ disconnect s
|
||||||
app $ putStrLn "Logged in"
|
app $ putStrLn "Logged in"
|
||||||
secret @@ s <- readSecret s
|
secret # s <- readSecret s
|
||||||
app $ putStrLn ("Secret: " ++ secret)
|
app $ putStrLn ("Secret: " ++ secret)
|
||||||
s <- logout s
|
s <- logout s
|
||||||
app $ putStrLn "Logged out"
|
app $ putStrLn "Logged out"
|
||||||
|
@ -25,8 +25,8 @@ login : (1 s : Store LoggedOut) -> (password : String) ->
|
|||||||
Res Bool (\ok => Store (if ok then LoggedIn else LoggedOut))
|
Res Bool (\ok => Store (if ok then LoggedIn else LoggedOut))
|
||||||
login (MkStore secret) password
|
login (MkStore secret) password
|
||||||
= if password == "Mornington Crescent"
|
= if password == "Mornington Crescent"
|
||||||
then True @@ MkStore secret
|
then True # MkStore secret
|
||||||
else False @@ MkStore secret
|
else False # MkStore secret
|
||||||
|
|
||||||
logout : (1 s : Store LoggedIn) -> Store LoggedOut
|
logout : (1 s : Store LoggedIn) -> Store LoggedOut
|
||||||
logout (MkStore secret) = MkStore secret
|
logout (MkStore secret) = MkStore secret
|
||||||
@ -37,9 +37,9 @@ storeProg
|
|||||||
do putStr "Password: "
|
do putStr "Password: "
|
||||||
password <- Console.getStr
|
password <- Console.getStr
|
||||||
connect $ \s =>
|
connect $ \s =>
|
||||||
do let True @@ s = login s password
|
do let True # s = login s password
|
||||||
| False @@ s => do putStrLn "Incorrect password"
|
| False # s => do putStrLn "Incorrect password"
|
||||||
disconnect s
|
disconnect s
|
||||||
putStrLn "Door opened"
|
putStrLn "Door opened"
|
||||||
let s = logout s
|
let s = logout s
|
||||||
putStrLn "Door closed"
|
putStrLn "Door closed"
|
||||||
|
@ -12,4 +12,4 @@ Error during reflection: Still not trying
|
|||||||
refprims.idr:43:10--45:1:While processing right hand side of dummy3 at refprims.idr:43:1--45:1:
|
refprims.idr:43:10--45:1:While processing right hand side of dummy3 at refprims.idr:43:1--45:1:
|
||||||
Error during reflection: Undefined name
|
Error during reflection: Undefined name
|
||||||
refprims.idr:46:10--48:1:While processing right hand side of dummy4 at refprims.idr:46:1--48:1:
|
refprims.idr:46:10--48:1:While processing right hand side of dummy4 at refprims.idr:46:1--48:1:
|
||||||
Error during reflection: failed after generating Main.{plus:6078}
|
Error during reflection: failed after generating Main.{plus:6087}
|
||||||
|
Loading…
Reference in New Issue
Block a user