├── .gitignore ├── INSTALL.txt ├── LICENSE ├── README.md ├── READMEconsolePrograms.txt ├── docs ├── Agda.Builtin.Bool.html ├── Agda.Builtin.Char.html ├── Agda.Builtin.Coinduction.html ├── Agda.Builtin.Equality.html ├── Agda.Builtin.List.html ├── Agda.Builtin.Nat.html ├── Agda.Builtin.Size.html ├── Agda.Builtin.String.html ├── Agda.Builtin.TrustMe.html ├── Agda.Builtin.Unit.html ├── Agda.Primitive.html ├── Agda.css ├── Algebra.FunctionProperties.Core.html ├── Algebra.FunctionProperties.html ├── Algebra.Morphism.html ├── Algebra.Operations.html ├── Algebra.Properties.AbelianGroup.html ├── Algebra.Properties.BooleanAlgebra.Expression.html ├── Algebra.Properties.BooleanAlgebra.html ├── Algebra.Properties.DistributiveLattice.html ├── Algebra.Properties.Group.html ├── Algebra.Properties.Lattice.html ├── Algebra.Properties.Ring.html ├── Algebra.RingSolver.AlmostCommutativeRing.html ├── Algebra.RingSolver.Lemmas.html ├── Algebra.RingSolver.Simple.html ├── Algebra.RingSolver.html ├── Algebra.Structures.html ├── Algebra.html ├── Category.Applicative.Indexed.html ├── Category.Applicative.html ├── Category.Functor.Identity.html ├── Category.Functor.html ├── Category.Monad.Identity.html ├── Category.Monad.Indexed.html ├── Category.Monad.html ├── Coinduction.html ├── Data.Bool.Base.html ├── Data.Bool.Properties.html ├── Data.Bool.html ├── Data.BoundedVec.Inefficient.html ├── Data.Char.Base.html ├── Data.Char.Core.html ├── Data.Char.html ├── Data.Colist.html ├── Data.Conat.html ├── Data.Empty.html ├── Data.Fin.Dec.html ├── Data.Fin.Properties.html ├── Data.Fin.Subset.Properties.html ├── Data.Fin.Subset.html ├── Data.Fin.html ├── Data.List.Any.html ├── Data.List.Base.html ├── Data.List.NonEmpty.html ├── Data.List.html ├── Data.Maybe.Base.html ├── Data.Maybe.html ├── Data.Nat.Base.html ├── Data.Nat.Properties.Simple.html ├── Data.Nat.Properties.html ├── Data.Nat.html ├── Data.Plus.html ├── Data.Product.html ├── Data.String.Base.html ├── Data.String.html ├── Data.Sum.html ├── Data.Unit.Base.html ├── Data.Unit.NonEta.html ├── Data.Unit.html ├── Data.Vec.Equality.html ├── Data.Vec.N-ary.html ├── Data.Vec.Properties.html ├── Data.Vec.html ├── Function.Bijection.html ├── Function.Equality.html ├── Function.Equivalence.html ├── Function.Injection.html ├── Function.Inverse.html ├── Function.LeftInverse.html ├── Function.Related.html ├── Function.Surjection.html ├── Function.html ├── Level.html ├── NativeIO.html ├── Relation.Binary.Consequences.Core.html ├── Relation.Binary.Consequences.html ├── Relation.Binary.Core.html ├── Relation.Binary.EqReasoning.html ├── Relation.Binary.HeterogeneousEquality.Core.html ├── Relation.Binary.HeterogeneousEquality.html ├── Relation.Binary.Indexed.Core.html ├── Relation.Binary.Indexed.html ├── Relation.Binary.InducedPreorders.html ├── Relation.Binary.Lattice.html ├── Relation.Binary.List.Pointwise.html ├── Relation.Binary.List.StrictLex.html ├── Relation.Binary.On.html ├── Relation.Binary.PartialOrderReasoning.html ├── Relation.Binary.PreorderReasoning.html ├── Relation.Binary.PropositionalEquality.Core.html ├── Relation.Binary.PropositionalEquality.TrustMe.html ├── Relation.Binary.PropositionalEquality.html ├── Relation.Binary.Reflection.html ├── Relation.Binary.Vec.Pointwise.html ├── Relation.Binary.html ├── Relation.Nullary.Decidable.html ├── Relation.Nullary.Negation.html ├── Relation.Nullary.html ├── Relation.Unary.html ├── Size.html ├── SizedIO.Base.html ├── SizedIO.Console.html ├── SizedIO.Object.html ├── StateSizedIO.Base.html ├── StateSizedIO.IOObject.html ├── StateSizedIO.Object.html ├── StateSizedIO.RObject.html ├── StateSizedIO.cellStateDependent.html ├── StateSizedIO.writingOOsUsingIO.html ├── StateSizedIO.writingOOsUsingIOVers4ReaderMethods.html ├── Unit.html ├── examples.heap.ALL.html ├── examples.heap.correctnessLinkedList.html ├── examples.heap.correctnessLinkedListStep2.html ├── heap.library.html ├── heap.libraryBool.html ├── heap.libraryFin.html ├── heap.libraryMaybe.html ├── heap.libraryNat.html ├── src.heap.heapAsObject.html ├── src.heap.heapAsObjectBase.html ├── src.heap.heapAsObjectExample.html ├── src.heap.heapAsObjectGeneric.html ├── src.heap.heapAsObjectNativeHeap.html └── src.heap.worldModule.html ├── exampleCode ├── helloWorld.agda └── myExample.agda-lib ├── examples ├── DrawingProgram │ ├── IOExampleGraphicsDrawingProgram.agda │ ├── IOExampleGraphicsDrawingProgramWhenJust.agda │ ├── IOExampleGraphicsMovingText.agda │ └── Main.agda-lib ├── Sized │ ├── ConsoleExample.agda │ ├── CounterCell.agda │ ├── Parrot.agda │ ├── SelfRef.agda │ ├── SimpleCell.agda │ └── StatefulCell.agda ├── SizedPoly │ ├── Console.agda │ └── SimpleCell.agda ├── StateSized │ ├── GUI │ │ ├── BitMaps.agda │ │ ├── ShipBitMap.agda │ │ ├── SpaceShipAdvanced.agda │ │ ├── SpaceShipCell.agda │ │ ├── SpaceShipExtendedExample.agda │ │ ├── SpaceShipSimpleVar.agda │ │ └── ship.ico │ ├── StackStateDependent.agda │ └── cellStateDependent.agda ├── UnSized │ ├── Console.agda │ ├── Parrot.agda │ ├── SelfRef.agda │ └── SimpleCell.agda ├── consoleExamples │ ├── helloWorld.agda │ ├── passwordCheckSimple.agda │ └── simpleLoop.agda ├── examplesPaperJFP │ ├── BasicIO.agda │ ├── CatNonTerm.agda │ ├── CatTerm.agda │ ├── Coalgebra.agda │ ├── Colists.agda │ ├── Collatz.agda │ ├── Console.agda │ ├── ConsoleInterface.agda │ ├── CounterCell.agda │ ├── Equality.agda │ ├── ExampleDrawingProgram.agda │ ├── IOExampleConsole.agda │ ├── IOGraphicsLib.agda │ ├── NativeIOSafe.agda │ ├── Object.agda │ ├── Sized.agda │ ├── SpaceShipAdvanced.agda │ ├── SpaceShipCell.agda │ ├── SpaceShipSimpleVar.agda │ ├── StackBisim.agda │ ├── StateDependentIO.agda │ ├── StatefulObject.agda │ ├── VariableList.agda │ ├── VariableListForDispatchOnly.agda │ ├── WxGraphicsLib.agda │ ├── agdaCodeBrady.agda │ ├── exampleFinFun.agda │ ├── finn.agda │ ├── loadAllOOAgdaFilesAsInLibrary.agda │ ├── loadAllOOAgdaPart1.agda │ ├── loadAllOOAgdaPart2.agda │ ├── safeFibStackMachineObjectOriented.agda │ └── triangleRightOperator.agda ├── java │ ├── Cell.java │ ├── CounterCell.java │ ├── Makefile │ ├── Parrot.java │ ├── SelfRefExample.java │ ├── SimpleCell.java │ ├── Stack │ │ ├── README │ │ ├── Stack.java │ │ └── StackLinkedList.java │ └── StatsCell.java └── ooAgdaExamples.agda-lib ├── presentationsAndExampleCode └── agdaImplementorsMeetingGlasgow22April2016AntonSetzer │ ├── NativeIOSafe.agda │ ├── StateSized │ └── GUI │ │ └── ship.ico │ ├── drawingProgram.agda │ ├── interactiveProgramsAgda.agda │ ├── interactiveProgramsAgda.old.agda │ ├── interactiveProgramsAgdaUnsized.agda │ ├── interfaceExtensionAndDelegation.agda │ ├── main.agda │ ├── objectOrientedGui.agda │ ├── objectsInAgda.agda │ ├── objectsInAgdaUnsized.agda │ ├── ship.ico │ ├── spaceShipAdvanced.agda │ ├── spaceShipCell.agda │ ├── spaceShipSimpleVar.agda │ ├── stateDependentIO.agda │ ├── stateDependentObjects.agda │ └── talkGlasgow.agda-lib └── src ├── ConsoleLib.agda ├── NativeIO.agda ├── NativeInt.agda ├── NativePolyIO.agda ├── SizedIO ├── Base.agda ├── Console.agda ├── ConsoleObject.agda ├── IOGraphicsLib.agda ├── IOObject.agda ├── Object.agda ├── coIOIO.agda └── coIOIOObject.agda ├── SizedPolyIO ├── Base.agda ├── Console.agda ├── ConsoleObject.agda ├── IOObject.agda └── Object.agda ├── StateSizedIO ├── Base.agda ├── Base.agda.README.txt ├── GUI │ ├── BaseStateDependent.agda │ ├── VariableList.agda │ ├── WxBindingsFFI.agda │ └── WxGraphicsLib.agda ├── IOObject.agda └── Object.agda ├── UnSizedIO ├── Base.agda ├── Console.agda ├── ConsoleObject.agda ├── IOObject.agda └── Object.agda ├── Unit.agda ├── lib ├── libraryDec.agda └── libraryString.agda └── ooAgda.agda-lib /.gitignore: -------------------------------------------------------------------------------- 1 | # Ignore Emacs backup files 2 | 3 | *~ 4 | \#* 5 | .\#* 6 | *~ 7 | \#*# 8 | \.#Makefile 9 | .[#]*[#] 10 | 11 | # Ignore Agda/Haskell/Java compiled files 12 | 13 | MAlonzo/ 14 | *.agdai 15 | *.hi 16 | *.o 17 | *.class 18 | *.pyg 19 | 20 | 21 | # Ignore LaTeX generated files 22 | 23 | *.log 24 | *.dvi 25 | *.ps 26 | *.blg 27 | *.bbl 28 | *.aux 29 | *_aux 30 | *.toc 31 | *.ptb 32 | *.nav 33 | *.out 34 | *.snm 35 | *.vrb 36 | 37 | # Misc 38 | 39 | *.rtf 40 | *.backup 41 | 42 | # executables 43 | IOExampleGraphicsDrawingProgram 44 | SpaceShipAdvanced 45 | SpaceShipExtendedExample 46 | examples/helloWorld 47 | examples/passwordCheckSimple -------------------------------------------------------------------------------- /INSTALL.txt: -------------------------------------------------------------------------------- 1 | INSTALL: 2 | 3 | This library contains the code as in the paper 4 | Andreas Abel, Stephan Adelsberger, Anton Setzer: Interactive Programming in 5 | Agda - Objects and Graphical User Interfaces 6 | Journal of Functional Programming, 27, 2017. 7 | doi: http://dx.doi.org/10.1017/S0956796816000319 8 | author's copy: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.pdf 9 | bibtex: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.bib 10 | 11 | The code works with the most recent version of Agda version 2.5 12 | (master branch from github, commit 13 | "useefd12c778e095ee783eeb1f845fc1938ea360204"). 14 | 15 | 16 | The current agda code uses the haskell graphic library "wxhaskell" and 17 | implements a small game, where a space ship can be moved with the arrow keys. 18 | 19 | 20 | Please note: file path are set for linux/unix, for Windows adapt the file 21 | "./examples/StateSized/GUI/ShipBitMap.agda". 22 | 23 | 24 | Details: 25 | 26 | - This version has been tested with Agda version 2.6.2-40e78a7 27 | 28 | - Install agda from git: 29 | 30 | git clone https://github.com/agda/agda.git 31 | 32 | sudo cabal update 33 | 34 | in the created directory 35 | 36 | sudo cabal install --global 37 | agda-mode setup 38 | 39 | - Install Agda standard library version (latest master branch version, 40 | tested with commit "545b94b81e4eb1d5a8b2b9f1083b12de5422d786") 41 | from https://github.com/agda/agda-stdlib 42 | 43 | git clone https://github.com/agda/agda-stdlib 44 | 45 | (adaption of library files see below) 46 | 47 | - If you want to write GUI programs such as the space ship example 48 | you need to install wxWidgets and wxHaskell as follows, 49 | otherwise (if you only want to work with objects in Agda or with console IO) 50 | you can continue directly with "Set up the agda library files" below 51 | 52 | 53 | - Install wxWidgets version 3.0.2 54 | https://www.wxwidgets.org/ 55 | Follow the installation instructions in for the latest stable version 56 | ("Latest Stable Release: 3.0.2" not the development version) 57 | 58 | Follow the instructions in 59 | 60 | https://github.com/wxWidgets/wxWidgets/blob/v3.0.2/docs/readme.txt 61 | 62 | and use the installation instructions for gtk, 63 | with instructions in the downloaded version 64 | at wxWidgets-3.1.0/docs/gtk/install.txt 65 | (easiest to follow instructions "The simplest case") 66 | 67 | - Install wxHaskell from github 68 | https://github.com/wxHaskell/wxHaskell 69 | (latest master branch version, tested with commit 70 | "db2a571300b5810a7dadfe117744e5b9389e0c9b"") 71 | 72 | Best is to clone 73 | https://github.com/wxHaskell/wxHaskell 74 | by using 75 | git clone https://github.com/wxHaskell/wxHaskell 76 | 77 | Then follow instructions at 78 | install.txt 79 | 80 | possibly replace there 81 | cabal install 82 | by 83 | sudo cabal install 84 | or 85 | sudo cabal install --global 86 | or 87 | cabal install --global 88 | 89 | 90 | 91 | 92 | - Set up the agda library files: 93 | 94 | In "~/.agda/libraries" add the following three lines: 95 | ----------- 96 | /agda-stdlib/standard-library.agda-lib 97 | /ooAgda/src/ooAgda.agda-lib 98 | /ooAgda/examples/ooAgdaExamples.agda-lib 99 | ---------- 100 | 101 | with the adapted to where your version of agda-stdlib 102 | (Agda standard lib) 103 | and of ooAgda is located 104 | 105 | - For using the libraries in your code add a library file 106 | containing as dependencies standard-library, ooAgdaExamples 107 | into your code. 108 | 109 | The helloworld example with an agda library file myExample.agda-lib 110 | can be found in subdirectory exampleCode 111 | 112 | 113 | For more information on the library files, 114 | please refer to the changelog of Agda "2.5.1": 115 | https://github.com/agda/agda/blob/master/CHANGELOG 116 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Copyright (c) 2009, 2016 Andreas Abel, Stephan Adelsberger, Shenggang Chen, 2 | Anton Setzer 3 | 4 | 5 | 6 | Permission is hereby granted, free of charge, to any person obtaining 7 | a copy of this software and associated documentation files (the 8 | "Software"), to deal in the Software without restriction, including 9 | without limitation the rights to use, copy, modify, merge, publish, 10 | distribute, sublicense, and/or sell copies of the Software, and to 11 | permit persons to whom the Software is furnished to do so, subject to 12 | the following conditions: 13 | 14 | The above copyright notice and this permission notice shall be 15 | included in all copies or substantial portions of the Software. 16 | 17 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, 18 | EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF 19 | MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. 20 | IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY 21 | CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, 22 | TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION WITH THE 23 | SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. 24 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Objects and Graphical User Interfaces in Agda 2 | 3 | This library contains the code as in the paper 4 | * Andreas Abel, Stephan Adelsberger, Anton Setzer: Interactive Programming in 5 | Agda - Objects and Graphical User Interfaces 6 | 7 | * Journal of Functional Programming, 27, 2017. 8 | * [doi 10.1145/2976022.2976032](http://dx.doi.org/10.1017/S0956796816000319) 9 | * [Author's copy](http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.pdf) 10 | * [BibTex](http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.bib) 11 | 12 | ## Installation 13 | * see INSTALL.txt 14 | 15 | ## Example code from the ooAgda paper 16 | 17 | * Code from the paper can be loaded using files 18 | * examples/examplesPaperJFP/loadAllOOAgdaPart1.agda 19 | * examples/examplesPaperJFP/loadAllOOAgdaPart2.agda 20 | * examplesPaperJFP/loadAllOOAgdaFilesAsInLibrary.agda 21 | 22 | ## Hello World Example 23 | * see exampleCode -------------------------------------------------------------------------------- /READMEconsolePrograms.txt: -------------------------------------------------------------------------------- 1 | Examples of console IO programs can be found in 2 | examples/consoleExamples/ 3 | 4 | The main library needed is in 5 | src/ConsoleLib.agda 6 | 7 | Note that there is no need to install wxWidgets and wxHaskell if you are 8 | only interested in console programs. 9 | 10 | For compiling program use Agda-> Compile and type in as backend GHC. 11 | If compiling the examples in examples/consoleExamples/ the executable programs 12 | will be in the directory examples 13 | -------------------------------------------------------------------------------- /docs/Agda.Builtin.Size.html: -------------------------------------------------------------------------------- 1 | 2 | Agda.Builtin.Size
{-# OPTIONS --without-K #-}
 28 | 
 29 | module Agda.Builtin.Size where
 42 | 
 43 | {-# BUILTIN SIZEUNIV SizeU #-}
 60 | {-# BUILTIN SIZE     Size   #-}
 77 | {-# BUILTIN SIZELT   Size<_ #-}
 94 | {-# BUILTIN SIZESUC  ↑_     #-}
111 | {-# BUILTIN SIZEINF  ω      #-}
128 | {-# BUILTIN SIZEMAX  _⊔ˢ_  #-}
145 | 
150 | -------------------------------------------------------------------------------- /docs/Agda.Builtin.TrustMe.html: -------------------------------------------------------------------------------- 1 | 2 | Agda.Builtin.TrustMe
{-# OPTIONS --without-K #-}
 28 | 
 29 | module Agda.Builtin.TrustMe where
 42 | 
 43 | open import Agda.Builtin.Equality
 56 | 
 57 | primitive primTrustMe :  {a} {A : Set a} {x y : A}  x  y
138 | 
143 | -------------------------------------------------------------------------------- /docs/Agda.Builtin.Unit.html: -------------------------------------------------------------------------------- 1 | 2 | Agda.Builtin.Unit
{-# OPTIONS --without-K #-}
 28 | 
 29 | module Agda.Builtin.Unit where
 42 | 
 43 | record  : Set where
 64 |   instance constructor tt
 77 | 
 78 | {-# BUILTIN UNIT  #-}
 95 | {-# COMPILE GHC  = data () (()) #-}
112 | 
117 | -------------------------------------------------------------------------------- /docs/Agda.css: -------------------------------------------------------------------------------- 1 | /* Aspects. */ 2 | .Comment { color: #B22222 } 3 | .Keyword { color: #CD6600 } 4 | .String { color: #B22222 } 5 | .Number { color: #A020F0 } 6 | .Symbol { color: #404040 } 7 | .PrimitiveType { color: #0000CD } 8 | .Operator {} 9 | 10 | /* NameKinds. */ 11 | .Bound { color: black } 12 | .InductiveConstructor { color: #008B00 } 13 | .CoinductiveConstructor { color: #8B7500 } 14 | .Datatype { color: #0000CD } 15 | .Field { color: #EE1289 } 16 | .Function { color: #0000CD } 17 | .Module { color: #A020F0 } 18 | .Postulate { color: #0000CD } 19 | .Primitive { color: #0000CD } 20 | .Record { color: #0000CD } 21 | 22 | /* OtherAspects. */ 23 | .DottedPattern {} 24 | .UnsolvedMeta { color: black; background: yellow } 25 | .UnsolvedConstraint { color: black; background: yellow } 26 | .TerminationProblem { color: black; background: #FFA07A } 27 | .IncompletePattern { color: black; background: #F5DEB3 } 28 | .Error { color: red; text-decoration: underline } 29 | .TypeChecks { color: black; background: #ADD8E6 } 30 | 31 | /* Standard attributes. */ 32 | a { text-decoration: none } 33 | a[href]:hover { background-color: #B4EEB4 } 34 | -------------------------------------------------------------------------------- /docs/Data.Char.Core.html: -------------------------------------------------------------------------------- 1 | 2 | Data.Char.Core
------------------------------------------------------------------------
16 | -- The Agda standard library
21 | --
26 | -- Basic definitions for Characters
31 | ------------------------------------------------------------------------
36 | 
37 | module Data.Char.Core where
50 | 
51 | ------------------------------------------------------------------------
56 | -- The type
61 | 
62 | open import Agda.Builtin.Char public using (Char)
91 | 
96 | -------------------------------------------------------------------------------- /docs/Data.Unit.Base.html: -------------------------------------------------------------------------------- 1 | 2 | Data.Unit.Base
------------------------------------------------------------------------
 16 | -- The Agda standard library
 21 | --
 26 | -- The unit type and the total relation on unit
 31 | ------------------------------------------------------------------------
 36 | module Data.Unit.Base where
 49 | 
 50 | ------------------------------------------------------------------------
 55 | -- A unit type defined as a record type
 60 | 
 61 | -- Note that the name of this type is "\top", not T.
 66 | 
 67 | open import Agda.Builtin.Unit public using (; tt)
102 | 
103 | record _≤_ (x y : ) : Set where
144 | 
149 | -------------------------------------------------------------------------------- /docs/Unit.html: -------------------------------------------------------------------------------- 1 | 2 | Unit

14 | module Unit where
27 | 
28 | record Unit : Set where
49 |   constructor unit
58 | 
59 | {-# COMPILED_DATA Unit () () #-}
76 | 
81 | -------------------------------------------------------------------------------- /exampleCode/helloWorld.agda: -------------------------------------------------------------------------------- 1 | module helloWorld where 2 | 3 | open import ConsoleLib 4 | 5 | main : ConsoleProg 6 | main = run (WriteString "Hello World") 7 | -------------------------------------------------------------------------------- /exampleCode/myExample.agda-lib: -------------------------------------------------------------------------------- 1 | name: myExample 2 | depend: standard-library, ooAgdaExamples 3 | include: . 4 | -------------------------------------------------------------------------------- /examples/DrawingProgram/IOExampleGraphicsDrawingProgram.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module IOExampleGraphicsDrawingProgram where 4 | 5 | 6 | open import Data.Bool.Base 7 | open import Data.Char.Base renaming (primCharEquality to charEquality) 8 | open import Data.Nat.Base hiding (_≟_;_⊓_; _+_; _*_) 9 | open import Data.List.Base hiding (_++_) 10 | open import Data.Integer.Base hiding (suc) 11 | open import Data.String.Base 12 | open import Data.Maybe.Base 13 | 14 | open import Function 15 | 16 | open import SizedIO.Base 17 | open import SizedIO.IOGraphicsLib 18 | 19 | open import NativeInt --PrimTypeHelpers 20 | open import NativeIO 21 | 22 | integerLess : ℤ → ℤ → Bool 23 | integerLess x y with ∣(y - (x ⊓ y))∣ 24 | ... | zero = true 25 | ... | z = false 26 | 27 | line : Point → Point → Graphic 28 | line p newpoint = withColor red (polygon 29 | (nativePoint x y 30 | ∷ nativePoint a b 31 | ∷ nativePoint (a + xAddition) (b + yAddition) 32 | ∷ nativePoint (x + xAddition) (y + yAddition) 33 | ∷ [] ) ) 34 | where 35 | x = nativeProj1Point p 36 | y = nativeProj2Point p 37 | a = nativeProj1Point newpoint 38 | b = nativeProj2Point newpoint 39 | 40 | diffx = + ∣ (a - x) ∣ 41 | diffy = + ∣ (b - y) ∣ 42 | 43 | diffx*3 = diffx * (+ 3) 44 | diffy*3 = diffy * (+ 3) 45 | 46 | condition = (integerLess diffx diffy*3) ∧ (integerLess diffy diffx*3) 47 | 48 | xAddition = if condition then + 2 else + 1 49 | yAddition = if condition then + 2 else + 1 50 | 51 | 52 | 53 | State = Maybe Point 54 | 55 | loop : ∀{i} → Window → State → IOGraphics i Unit 56 | force (loop w s) = exec' (getWindowEvent w) 57 | λ{ (Key c t) → if charEquality c 'x' then (exec (closeWindow w) return) 58 | else loop w s 59 | ; (MouseMove p₂) → case s of 60 | λ{ nothing → loop w (just p₂) 61 | ; (just p₁) → exec (drawInWindow w (line p₁ p₂)) λ _ → 62 | loop w (just p₂) 63 | } 64 | ; _ → loop w s 65 | } 66 | 67 | 68 | {- 69 | loop : ∀{i} → Window → State → IOGraphics i Unit 70 | force (loop w s) = exec' (getWindowEvent w) 71 | λ{ (Key c t) → if charEquality c 'x' then return _ else loop w s 72 | ; (MouseMove p₂) → case s of 73 | λ{ nothing → loop w (just p₂) 74 | ; (just p₁) → exec (drawInWindow w (line p₁ p₂)) λ _ → 75 | loop w (just p₂) 76 | } 77 | ; _ → loop w s 78 | } 79 | -} 80 | 81 | {- 82 | mutual 83 | loop : ∀{i} → Window → State → IOGraphics i Unit 84 | force (loop w s) = exec' (getWindowEvent w) (λ e → aux w e s) 85 | 86 | aux : ∀{i} → Window → Event → State → IOGraphics i Unit 87 | aux w (Key c t) s = if charEquality c 'x' then (exec (closeWindow w) λ _ → return _) 88 | else loop w s 89 | aux w (MouseMove p₂) (just p₁) = exec (drawInWindow w (line p₁ p₂)) (λ _ → 90 | loop w (just p₂)) 91 | aux w (MouseMove p₂) nothing = loop w (just p₂) 92 | aux w _ s = loop w s 93 | -} 94 | 95 | {- 96 | tailrecursion discussion 97 | 98 | mutual 99 | loop : ∀{i} → Window → State → IOGraphics i Unit 100 | command (force (loop w s)) = (getWindowEvent w) 101 | response (force (loop w s)) e = aux w e s 102 | 103 | aux : ∀{i} → Window → Event → State → IOGraphics i Unit 104 | command (aux w (MouseMove p₂) (just p₁)) = (drawInWindow w (line p₁ p₂)) 105 | response (aux w (MouseMove p₂) (just p₁)) _ = loop w (just p₂)) 106 | aux w (MouseMove p₂) nothing = loop w (just p₂) 107 | aux w _ s = loop w s 108 | 109 | -} 110 | 111 | 112 | {- 113 | λ{ (Key c t) → if charEquality c 'x' then 114 | (exec ? λ _ → return _) else loop w s 115 | ; (MouseMove p₂) → case s of 116 | λ{ nothing → loop w (just p₂) 117 | ; (just p₁) → exec (drawInWindow w (line p₁ p₂)) λ _ → 118 | loop w (just p₂) 119 | } 120 | ; _ → loop w s 121 | } 122 | -} 123 | 124 | 125 | 126 | myProgram : ∀{i} → IOGraphics i Unit 127 | myProgram = 128 | exec (openWindow "Drawing Program" nothing 129 | 1000 1000 nativeDrawGraphic nothing) λ window → 130 | loop window nothing 131 | 132 | 133 | main : NativeIO Unit 134 | main = nativeRunGraphics (translateIOGraphics myProgram) 135 | -------------------------------------------------------------------------------- /examples/DrawingProgram/IOExampleGraphicsDrawingProgramWhenJust.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module IOExampleGraphicsDrawingProgramWhenJust where 4 | 5 | 6 | open import Data.Bool.Base 7 | open import Data.Char.Base renaming (primCharEquality to charEquality) 8 | open import Data.Nat.Base hiding (_≟_;_⊓_; _+_; _*_) 9 | open import Data.List.Base hiding (_++_) 10 | open import Data.Integer.Base hiding (suc) 11 | open import Data.String.Base 12 | open import Data.Maybe.Base 13 | 14 | open import Function 15 | 16 | open import SizedIO.Base 17 | open import SizedIO.IOGraphicsLib 18 | 19 | open import NativeInt --PrimTypeHelpers 20 | open import NativeIO 21 | 22 | integerLess : ℤ → ℤ → Bool 23 | integerLess x y with ∣(y - (x ⊓ y))∣ 24 | ... | zero = true 25 | ... | z = false 26 | 27 | line : Point → Point → Graphic 28 | line p newpoint = withColor red (polygon 29 | (point x y 30 | ∷ point a b 31 | ∷ point (a + xAddition) (b + yAddition) 32 | ∷ point (x + xAddition) (y + yAddition) 33 | ∷ [] ) ) 34 | where 35 | x = proj1Point p 36 | y = proj2Point p 37 | a = proj1Point newpoint 38 | b = proj2Point newpoint 39 | 40 | diffx = + ∣ (a - x) ∣ 41 | diffy = + ∣ (b - y) ∣ 42 | 43 | diffx*3 = diffx * (+ 3) 44 | diffy*3 = diffy * (+ 3) 45 | 46 | condition = (integerLess diffx diffy*3) ∧ (integerLess diffy diffx*3) 47 | 48 | xAddition = if condition then + 2 else + 1 49 | yAddition = if condition then + 2 else + 1 50 | 51 | 52 | 53 | State = Maybe Point 54 | 55 | loop : ∀{i} → Window → State → IOGraphics i Unit 56 | force (loop w s) = exec' (getWindowEvent w) 57 | λ{ (Key c t) → if charEquality c 'x' then return _ else loop w s 58 | ; (MouseMove p₂) → whenJust s (λ p₁ → exec1 (drawInWindow w (line p₁ p₂))) >>= λ _ → 59 | loop w (just p₂) 60 | ; _ → loop w s 61 | } 62 | 63 | myProgram : ∀{i} → IOGraphics i Unit 64 | myProgram = 65 | exec (openWindow "Drawing Program" nothing 66 | (just (size (+ 1000) (+ 1000))) nativeDrawGraphic nothing) λ window → 67 | loop window nothing 68 | 69 | 70 | main : NativeIO Unit 71 | main = nativeRunGraphics (translateIOGraphics myProgram) 72 | -------------------------------------------------------------------------------- /examples/DrawingProgram/IOExampleGraphicsMovingText.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module IOExampleGraphicsMovingText where 4 | 5 | -- 6 | -- not yet updated for new SizedIO library 7 | -- not yet updated to change with IOInterface Record of Commands / Responses 8 | -- 9 | 10 | {- 11 | 12 | open import SizedIO.General 13 | open import SizedIO.IOGraphicsLib -- TODO 14 | 15 | open import Data.Bool.Base 16 | open import Data.Char.Base 17 | open import Data.Nat.Base 18 | open import Data.List.Base hiding (_++_) 19 | open import Data.Integer.Base hiding (_+_;suc) 20 | open import Data.String.Base 21 | open import Data.Maybe.Base 22 | 23 | open import Function 24 | open import Size 25 | 26 | width : ℕ 27 | width = 125 28 | 29 | height : ℕ 30 | height = 20 31 | 32 | myBox : Point → Graphic 33 | myBox p = withColor black (polygon ( point x y 34 | ∷ point (x +IntNat width) y 35 | ∷ point (x +IntNat width) (y +IntNat height) 36 | ∷ point x (y +IntNat height) 37 | ∷ [] ) ) 38 | where x = proj1Point p 39 | y = proj2Point p 40 | 41 | record State : Set where 42 | constructor state 43 | field oldpoint : Maybe Point 44 | newpoint : Maybe Point 45 | open State 46 | 47 | 48 | module _ (w : Window) where 49 | 50 | -- If there was a mouse movement since the last invokation of refresh 51 | -- overwrite old text and place text at new mouse position. 52 | refresh : ∀{i} → State → IOGraphics i State 53 | refresh s = case newpoint s of 54 | λ{ nothing → return s 55 | ; (just lastMousePoint) → 56 | {- case there was a new mouse movement. -} 57 | (whenJust (oldpoint s) λ p → 58 | exec1 (drawInWindow w (myBox p)) >> 59 | exec1 (drawInWindow w (withColor white (text lastMousePoint ("1st :some Text"))))) 60 | >> return (state (just lastMousePoint) nothing) 61 | } 62 | 63 | processEvent : ∀{i} → State → IOGraphics+ i (Maybe State) 64 | processEvent s = 65 | exec (getWindowEvent w) 66 | λ{ {- key event: check whether this was equal to 'x'; if yes, terminate, otherwise return to loop -} 67 | (Key c t) → 68 | if (c ==CharBool 'x') then return nothing else return (just s) 69 | 70 | {- mouse event: if this was the first ever event, ignore it, but 71 | record that the next mouse event is no longer the first one. 72 | If it was not the first ever event: set newpoint to the position of the mouse p -} 73 | ; (MouseMove p) → 74 | return (just record s{ newpoint = just p }) 75 | 76 | {- otherwise just go back to start -} 77 | ; _ → 78 | return (just s) 79 | } 80 | 81 | loop : ∀{i} → State → IOGraphics i Unit 82 | force (loop s) = processEvent s >>=+' -- maybe′ loop (return _) 83 | λ{ nothing → return _ 84 | ; (just s') → refresh s' >>= loop 85 | } 86 | 87 | myProgram : ∀{i} → IOGraphics i Unit 88 | force myProgram = 89 | exec (openWindowEx "Hello World" nothing (just (natSize 1000 1000)) nativeDrawGraphic nothing) λ window → 90 | loop window (state nothing nothing) 91 | 92 | main : NativeIO Unit 93 | main = nativeRunGraphics (translateIOGraphics myProgram) 94 | -} 95 | -------------------------------------------------------------------------------- /examples/DrawingProgram/Main.agda-lib: -------------------------------------------------------------------------------- 1 | name: Main 2 | depend: std-lib ionew 3 | include: . 4 | -------------------------------------------------------------------------------- /examples/Sized/ConsoleExample.agda: -------------------------------------------------------------------------------- 1 | module Sized.ConsoleExample where 2 | 3 | open import SizedIO.Base 4 | open import SizedIO.Console hiding (main) 5 | 6 | open import NativeIO 7 | 8 | myProgram : ∀{i} → IOConsole i Unit 9 | force myProgram = do' getLine λ line → 10 | do (putStrLn line) λ _ → 11 | do (putStrLn line) λ _ → 12 | myProgram 13 | 14 | 15 | main : NativeIO Unit 16 | main = translateIOConsole myProgram 17 | -------------------------------------------------------------------------------- /examples/Sized/CounterCell.agda: -------------------------------------------------------------------------------- 1 | module Sized.CounterCell where 2 | 3 | open import Data.Product 4 | open import Data.Nat.Base 5 | open import Data.Nat.Show 6 | open import Data.String.Base using (String; _++_) 7 | 8 | open import SizedIO.Object 9 | open import SizedIO.IOObject 10 | 11 | open import SizedIO.Base 12 | open import SizedIO.Console hiding (main) 13 | open import SizedIO.ConsoleObject 14 | 15 | open import NativeIO 16 | 17 | open import Sized.SimpleCell hiding (program; main) 18 | 19 | open import Size 20 | 21 | data CounterMethod A : Set where 22 | super : (m : CellMethod A) → CounterMethod A 23 | stats : CounterMethod A 24 | 25 | pattern getᶜ = super get 26 | pattern putᶜ x = super (put x) 27 | 28 | -- CounterResult : ∀{A} → 29 | 30 | counterI : (A : Set) → Interface 31 | Method (counterI A) = CounterMethod A 32 | Result (counterI A) (super m) = Result (cellJ A) m 33 | Result (counterI A) stats = Unit 34 | 35 | 36 | CounterC : (i : Size) → Set 37 | CounterC i = ConsoleObject i (counterI String) 38 | 39 | -- counterP is constructor for the consoleObject for interface counterI 40 | 41 | counterP : ∀{i} (c : CellC i) (ngets nputs : ℕ) → CounterC i 42 | method (counterP c ngets nputs) getᶜ = 43 | method c get >>= λ { (s , c') → 44 | return (s , counterP c' (1 + ngets) nputs) } 45 | method (counterP c ngets nputs) (putᶜ x) = 46 | method c (put x) >>= λ { (_ , c') → 47 | return (_ , counterP c' ngets (1 + nputs)) } 48 | method (counterP c ngets nputs) stats = 49 | exec (putStrLn ("Counted " 50 | ++ show ngets ++ " calls to get and " 51 | ++ show nputs ++ " calls to put.")) λ _ → 52 | return (_ , counterP c ngets nputs) 53 | 54 | program : String → IOConsole ∞ Unit 55 | program arg = 56 | let c₀ = counterP (cellP "Start") 0 0 in 57 | method c₀ getᶜ >>= λ{ (s , c₁) → 58 | exec1 (putStrLn s) >> 59 | method c₁ (putᶜ arg) >>= λ{ (_ , c₂) → 60 | method c₂ getᶜ >>= λ{ (s' , c₃) → 61 | exec1 (putStrLn s') >> 62 | method c₃ (putᶜ "Over!") >>= λ{ (_ , c₄) → 63 | method c₄ stats >>= λ{ (_ , c₅) → 64 | return _ }}}}} 65 | 66 | main : NativeIO Unit 67 | main = translateIOConsole (program "Hello") 68 | 69 | -- -} 70 | -- -} 71 | -- -} 72 | -- -} 73 | -- -} 74 | -- -} 75 | -- -} 76 | -------------------------------------------------------------------------------- /examples/Sized/Parrot.agda: -------------------------------------------------------------------------------- 1 | module Sized.Parrot where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import SizedIO.IOObject 7 | 8 | open import SizedIO.Base 9 | open import SizedIO.Console hiding (main) 10 | open import SizedIO.ConsoleObject 11 | 12 | open import NativeIO 13 | 14 | open import Sized.SimpleCell hiding (program; main) 15 | 16 | open import Size 17 | 18 | record Wrap A : Set where 19 | constructor wrap 20 | field unwrap : A 21 | 22 | parrotI = cellJ (Wrap String) 23 | 24 | ParrotC : (i : Size) → Set 25 | ParrotC i = ConsoleObject i parrotI 26 | 27 | -- but reusing cellC from SimpleCell, as interface is ident! 28 | 29 | -- class Parrot implements Cell { 30 | -- Cell cell; 31 | -- Parrot (Cell c) { cell = c; } 32 | -- public String get() { 33 | -- return "(" + cell.get() + ") is what parrot got"; 34 | -- } 35 | -- public void put (String s) { 36 | -- cell.put("parrot puts (" + s + ")"); 37 | -- } 38 | -- } 39 | 40 | -- parrotP is constructor for the consoleObject for interface (cellI String) 41 | 42 | parrotP : ∀{i} (c : CellC i) → ParrotC i 43 | (method (parrotP c) get) = 44 | method c get >>= λ { (s , c') → 45 | return (wrap ("(" ++ s ++ ") is what parrot got") , parrotP c' ) } 46 | (method (parrotP c) (put (wrap s))) = 47 | method c (put ("parrot puts (" ++ s ++ ")")) >>= λ { (_ , c') → 48 | return (_ , parrotP c') } 49 | 50 | 51 | 52 | -- public static void main (String[] args) { 53 | -- Parrot c = new Parrot(new SimpleCell("Start")); 54 | -- System.out.println(c.get()); 55 | -- c.put(args[1]); 56 | -- System.out.println(c.get()); 57 | -- } 58 | -- } 59 | 60 | program : String → IOConsole ∞ Unit 61 | program arg = 62 | let c₀ = parrotP (cellP "Start") in 63 | method c₀ get >>= λ{ (wrap s , c₁) → 64 | exec1 (putStrLn s) >> 65 | method c₁ (put (wrap arg)) >>= λ{ (_ , c₂) → 66 | method c₂ get >>= λ{ (wrap s' , c₃) → 67 | exec1 (putStrLn s') }}} 68 | 69 | main : NativeIO Unit 70 | main = translateIOConsole (program "hello") 71 | 72 | -- -} 73 | -------------------------------------------------------------------------------- /examples/Sized/SelfRef.agda: -------------------------------------------------------------------------------- 1 | module Sized.SelfRef where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import Size 7 | 8 | open import SizedIO.Object 9 | open import SizedIO.IOObject 10 | open import SizedIO.ConsoleObject 11 | 12 | open import SizedIO.Base 13 | open import SizedIO.Console hiding (main) 14 | 15 | open import NativeIO 16 | 17 | 18 | data AMethod : Set where 19 | print : AMethod 20 | m1 : AMethod 21 | m2 : AMethod 22 | 23 | AResult : (c : AMethod) → Set 24 | AResult _ = Unit 25 | 26 | aI : Interface 27 | Method aI = AMethod 28 | Result aI = AResult 29 | 30 | aC : (i : Size) → Set 31 | aC i = ConsoleObject i aI 32 | 33 | -- 34 | -- Self Referential: method 'm1' calls method 'm2' 35 | -- 36 | {-# NON_TERMINATING #-} 37 | aP : ∀{i} (s : String) → aC i 38 | method (aP s) print = 39 | exec1 (putStrLn s) >> 40 | return (_ , aP s) 41 | 42 | method (aP s) m1 = 43 | exec1 (putStrLn s) >> 44 | method (aP s) m2 >>= λ{ (_ , c₀) → 45 | return (_ , c₀) } 46 | method (aP s) m2 = 47 | return (_ , aP (s ++ "->m2")) 48 | 49 | 50 | program : String → IOConsole ∞ Unit 51 | program arg = 52 | let c₀ = aP ("start̄") in 53 | method c₀ m1 >>= λ{ (_ , c₁) → --- ===> m1 called, then m2 prints out text 54 | method c₁ print >>= λ{ (_ , c₂) → 55 | exec1 (putStrLn "end") }} 56 | 57 | main : NativeIO Unit 58 | main = translateIOConsole (program "") 59 | -------------------------------------------------------------------------------- /examples/Sized/SimpleCell.agda: -------------------------------------------------------------------------------- 1 | module Sized.SimpleCell where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import SizedIO.Object 7 | open import SizedIO.IOObject 8 | open import SizedIO.ConsoleObject 9 | 10 | open import SizedIO.Base 11 | open import SizedIO.Console hiding (main) 12 | 13 | open import NativeIO 14 | 15 | open import Size 16 | 17 | data CellMethod A : Set where 18 | get : CellMethod A 19 | put : A → CellMethod A 20 | 21 | CellResult : ∀{A} → CellMethod A → Set 22 | CellResult {A} get = A 23 | CellResult (put _) = Unit 24 | 25 | -- cellJ is the interface of the object of a simple cell 26 | cellJ : (A : Set) → Interface 27 | Method (cellJ A) = CellMethod A 28 | Result (cellJ A) m = CellResult m 29 | 30 | cell : {A : Set} → A → Object (cellJ A) 31 | objectMethod (cell a) get = a , (cell a) 32 | objectMethod (cell a) (put a') = _ , (cell a') 33 | 34 | -- cellC is the type of consoleObjects with interface (cellJ String) 35 | CellC : (i : Size) → Set 36 | CellC i = ConsoleObject i (cellJ String) 37 | 38 | 39 | -- cellO is a program for a simple cell which 40 | -- when get is called writes "getting s" for the string s of the object 41 | -- and when putting s writes "putting s" for the string 42 | 43 | -- cellP is constructor for the consoleObject for interface (cellJ String) 44 | cellP : ∀{i} (s : String) → CellC i 45 | force (method (cellP s) get) = 46 | exec' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 47 | return (s , cellP s) 48 | force (method (cellP s) (put x)) = 49 | exec' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 50 | return (_ , (cellP x)) 51 | 52 | -- Program is another program 53 | program : String → IOConsole ∞ Unit 54 | program arg = 55 | let c₀ = cellP "Start" in 56 | method c₀ get >>= λ{ (s , c₁) → 57 | exec1 (putStrLn s) >> 58 | method c₁ (put arg) >>= λ{ (_ , c₂) → 59 | method c₂ get >>= λ{ (s' , c₃) → 60 | exec1 (putStrLn s') 61 | }}} 62 | 63 | main : NativeIO Unit 64 | main = translateIOConsole (program "hello") 65 | -------------------------------------------------------------------------------- /examples/Sized/StatefulCell.agda: -------------------------------------------------------------------------------- 1 | module Sized.StatefulCell where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import SizedIO.Object 7 | open import SizedIO.IOObject 8 | open import SizedIO.ConsoleObject 9 | 10 | open import SizedIO.Console hiding (main) 11 | 12 | open import NativeIO 13 | 14 | open import Size 15 | 16 | --open import StateSizedIO.Base 17 | 18 | data CellState# : Set where 19 | empty full : CellState# 20 | 21 | data CellMethodEmpty A : Set where 22 | put : A → CellMethodEmpty A 23 | 24 | data CellMethodFull A : Set where 25 | get : CellMethodFull A 26 | put : A → CellMethodFull A 27 | 28 | CellMethod# : ∀{A} → CellState# → Set 29 | CellMethod#{A} empty = CellMethodEmpty A 30 | CellMethod#{A} full = CellMethodFull A 31 | 32 | 33 | CellResultFull : ∀{A} → CellMethodFull A → Set 34 | CellResultFull {A} get = A 35 | CellResultFull (put _) = Unit 36 | 37 | CellResultEmpty : ∀{A} → CellMethodEmpty A → Set 38 | CellResultEmpty (put _) = Unit 39 | 40 | 41 | CellResult# : ∀{A} → (s : CellState#) → CellMethod#{A} s → Set 42 | CellResult#{A} empty = CellResultEmpty{A} 43 | CellResult#{A} full = CellResultFull{A} 44 | 45 | 46 | n# : ∀{A} → (s : CellState#) → (c : CellMethod#{A} s) → (CellResult# s c) → CellState# 47 | n# empty (put x) unit = full 48 | n# full get z = full 49 | n# full (put x) unit = full 50 | -------------------------------------------------------------------------------- /examples/SizedPoly/Console.agda: -------------------------------------------------------------------------------- 1 | module SizedPoly.Console where 2 | 3 | open import SizedPolyIO.Base 4 | open import SizedPolyIO.Console hiding (main) 5 | 6 | open import NativePolyIO 7 | 8 | open import Level using () renaming (zero to lzero) 9 | 10 | 11 | {-# TERMINATING #-} 12 | myProgram : ∀{i} → IOConsole i (Unit {lzero}) 13 | myProgram = exec getLine λ line → 14 | exec (putStrLn line) λ _ → 15 | exec (putStrLn line) λ _ → 16 | myProgram 17 | 18 | 19 | main : NativeIO Unit 20 | main = translateIOConsole myProgram 21 | -------------------------------------------------------------------------------- /examples/SizedPoly/SimpleCell.agda: -------------------------------------------------------------------------------- 1 | module SizedPoly.SimpleCell where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import Size 7 | open import Level using (_⊔_) renaming (suc to lsuc; zero to lzero) 8 | 9 | open import SizedPolyIO.Object 10 | open import SizedPolyIO.IOObject 11 | open import SizedPolyIO.ConsoleObject 12 | 13 | open import SizedPolyIO.Base 14 | open import SizedPolyIO.Console hiding (main) 15 | 16 | open import NativePolyIO 17 | 18 | data CellMethod A : Set where 19 | get : CellMethod A 20 | put : A → CellMethod A 21 | 22 | CellResponse : ∀{A} → CellMethod A → Set 23 | CellResponse {A} get = A 24 | CellResponse (put _) = Unit 25 | 26 | -- cellI is the interface of the object of a simple cell 27 | cellI : (A : Set) → Interface lzero lzero 28 | Method (cellI A) = CellMethod A 29 | Result (cellI A) m = CellResponse m 30 | 31 | -- cellC is the type of consoleObjects with interface (cellI String) 32 | CellC : (i : Size) → Set 33 | CellC i = ConsoleObject i (cellI String) 34 | 35 | -- cellP is constructor for the consoleObject for interface (cellI String) 36 | cellP : ∀{i} (s : String) → CellC i 37 | force (method (cellP s) get) = 38 | exec' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 39 | return (s , cellP s) 40 | force (method (cellP s) (put x)) = 41 | exec' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 42 | return (_ , (cellP x)) 43 | 44 | 45 | -- Program is another program 46 | program : String → IOConsole ∞ Unit 47 | program arg = 48 | let c₀ = cellP "Start" in 49 | method c₀ get >>= λ{ (s , c₁) → 50 | exec1 (putStrLn s) >> 51 | method c₁ (put arg) >>= λ{ (_ , c₂) → 52 | method c₂ get >>= λ{ (s' , c₃) → 53 | exec1 (putStrLn s') }}} 54 | 55 | main : NativeIO Unit 56 | main = translateIOConsole (program "hello") 57 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/BitMaps.agda: -------------------------------------------------------------------------------- 1 | module StateSized.GUI.BitMaps where 2 | 3 | open import StateSizedIO.GUI.WxBindingsFFI 4 | 5 | ship : Bitmap 6 | ship = bitmap "./StateSized/GUI/ship.ico" 7 | 8 | rock : Bitmap 9 | rock = bitmap "./StateSized/GUI/rock.ico" 10 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/ShipBitMap.agda: -------------------------------------------------------------------------------- 1 | module StateSized.GUI.ShipBitMap where 2 | 3 | open import StateSizedIO.GUI.WxBindingsFFI 4 | 5 | ship : Bitmap 6 | ship = bitmap "./StateSized/GUI/ship.ico" 7 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/SpaceShipAdvanced.agda: -------------------------------------------------------------------------------- 1 | module StateSized.GUI.SpaceShipAdvanced where 2 | 3 | 4 | open import SizedIO.Base 5 | open import StateSizedIO.GUI.BaseStateDependent 6 | 7 | 8 | open import Data.Bool.Base 9 | open import Data.List.Base 10 | open import Data.Integer 11 | 12 | 13 | open import Data.Product hiding (map) 14 | open import SizedIO.Object 15 | open import SizedIO.IOObject 16 | 17 | open import NativeIO 18 | open import Size 19 | 20 | open import StateSizedIO.GUI.WxBindingsFFI 21 | open import StateSizedIO.GUI.VariableList 22 | 23 | open import StateSizedIO.GUI.WxGraphicsLib 24 | 25 | open import StateSized.GUI.BitMaps 26 | 27 | data GraphicServerMethod : Set where 28 | onPaintM : DC → Rect → GraphicServerMethod 29 | moveSpaceShipM : Frame → GraphicServerMethod 30 | repaintM : Frame → GraphicServerMethod 31 | 32 | GraphicServerResult : GraphicServerMethod → Set 33 | GraphicServerResult _ = Unit 34 | 35 | GraphicServerInterface : Interface 36 | Method GraphicServerInterface = GraphicServerMethod 37 | Result GraphicServerInterface = GraphicServerResult 38 | 39 | GraphicServerObject : ∀{i} → Set 40 | GraphicServerObject {i} = IOObject GuiLev1Interface GraphicServerInterface i 41 | 42 | graphicServerObject : ∀{i} → ℤ → GraphicServerObject {i} 43 | method (graphicServerObject z) (onPaintM dc rect) = 44 | exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 45 | return (_ , graphicServerObject z) 46 | method (graphicServerObject z) (moveSpaceShipM fra) = 47 | return (_ , graphicServerObject (z + + 20)) 48 | method (graphicServerObject z) (repaintM fra) = 49 | exec (repaint fra) λ _ → 50 | return (_ , graphicServerObject z) 51 | 52 | VarType : Set 53 | VarType = GraphicServerObject {∞} 54 | 55 | varInit : VarType 56 | varInit = graphicServerObject (+ 150) 57 | 58 | 59 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 60 | onPaint obj dc rect = mapIO proj₂ (method obj (onPaintM dc rect)) 61 | 62 | 63 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 64 | moveSpaceShip fra obj = mapIO proj₂ (method obj (moveSpaceShipM fra)) 65 | 66 | 67 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 68 | callRepaint fra obj = mapIO proj₂ (method obj (repaintM fra)) 69 | 70 | 71 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 72 | program = execˢ (level1C makeFrame) λ fra → 73 | execˢ (level1C (makeButton fra)) λ bt → 74 | execˢ (level1C (addButton fra bt)) λ _ → 75 | execˢ (createVar varInit) λ _ → 76 | execˢ (setButtonHandler bt (moveSpaceShip fra 77 | ∷ [ callRepaint fra ])) λ _ → 78 | execˢ (setOnPaint fra [ onPaint ]) 79 | returnˢ 80 | 81 | main : NativeIO Unit 82 | main = start (translateLev2 program) 83 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/SpaceShipCell.agda: -------------------------------------------------------------------------------- 1 | module StateSized.GUI.SpaceShipCell where 2 | 3 | 4 | open import SizedIO.Base 5 | open import StateSizedIO.GUI.BaseStateDependent 6 | 7 | 8 | open import Data.Bool.Base 9 | open import Data.List.Base 10 | open import Data.Integer 11 | 12 | open import Data.Product hiding (map) 13 | open import SizedIO.Object 14 | open import SizedIO.IOObject 15 | 16 | open import NativeIO 17 | open import Sized.SimpleCell hiding (main; program) 18 | 19 | open import StateSizedIO.GUI.WxBindingsFFI 20 | open import StateSizedIO.GUI.VariableList 21 | 22 | open import StateSizedIO.GUI.WxGraphicsLib 23 | 24 | open import StateSized.GUI.BitMaps 25 | 26 | 27 | VarType = Object (cellJ ℤ) 28 | 29 | cellℤC : (z : ℤ ) → VarType 30 | objectMethod (cellℤC z) get = ( z , cellℤC z) 31 | objectMethod (cellℤC z) (put z') = (_ , cellℤC z') 32 | 33 | varInit : VarType 34 | varInit = cellℤC (+ 150) 35 | 36 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 37 | onPaint c dc rect = let (z , c₁) = objectMethod c get 38 | in exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 39 | return c₁ 40 | 41 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 42 | moveSpaceShip fra c = let (z , c₁) = objectMethod c get 43 | (_ , c₂) = objectMethod c₁ (put (z + + 20)) 44 | in return c₂ 45 | 46 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 47 | callRepaint fra c = exec (repaint fra) λ _ → 48 | return c 49 | 50 | 51 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 52 | program = execˢ (level1C makeFrame) λ fra → 53 | execˢ (level1C (makeButton fra)) λ bt → 54 | execˢ (level1C (addButton fra bt)) λ _ → 55 | execˢ (createVar varInit) λ _ → 56 | execˢ (setButtonHandler bt (moveSpaceShip fra 57 | ∷ [ callRepaint fra ])) λ _ → 58 | execˢ (setOnPaint fra [ onPaint ]) 59 | returnˢ 60 | 61 | main : NativeIO Unit 62 | main = start (translateLev2 program) 63 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/SpaceShipSimpleVar.agda: -------------------------------------------------------------------------------- 1 | module StateSized.GUI.SpaceShipSimpleVar where 2 | 3 | open import SizedIO.Base 4 | open import StateSizedIO.GUI.BaseStateDependent 5 | 6 | open import Data.Bool.Base 7 | open import Data.List.Base 8 | open import Data.Integer 9 | 10 | open import Data.Product hiding (map) 11 | open import SizedIO.Object 12 | open import SizedIO.IOObject 13 | 14 | open import NativeIO 15 | 16 | open import StateSizedIO.GUI.WxBindingsFFI 17 | open import StateSizedIO.GUI.VariableList 18 | 19 | open import StateSizedIO.GUI.WxGraphicsLib 20 | 21 | open import StateSized.GUI.BitMaps 22 | 23 | VarType : Set 24 | VarType = ℤ 25 | 26 | varInit : VarType 27 | varInit = (+ 150) 28 | 29 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 30 | onPaint z dc rect = exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 31 | return z 32 | 33 | 34 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 35 | moveSpaceShip fra z = return (z + + 20) 36 | 37 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 38 | callRepaint fra z = exec (repaint fra) λ _ → 39 | return z 40 | 41 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 42 | program = execˢ (level1C makeFrame) λ fra → 43 | execˢ (level1C (makeButton fra)) λ bt → 44 | execˢ (level1C (addButton fra bt)) λ _ → 45 | execˢ (createVar varInit) λ _ → 46 | execˢ (setButtonHandler bt (moveSpaceShip fra 47 | ∷ [ callRepaint fra ])) λ _ → 48 | execˢ (setOnPaint fra [ onPaint ]) 49 | returnˢ 50 | 51 | main : NativeIO Unit 52 | main = start (translateLev2 program) 53 | -------------------------------------------------------------------------------- /examples/StateSized/GUI/ship.ico: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/ooAgda/7cc45e0148a4a508d20ed67e791544c30fecd795/examples/StateSized/GUI/ship.ico -------------------------------------------------------------------------------- /examples/StateSized/StackStateDependent.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --postfix-projections #-} 2 | 3 | module StateSized.StackStateDependent where 4 | 5 | open import Data.Product 6 | open import Function 7 | open import Data.String.Base as Str 8 | open import Data.Nat.Base as N 9 | open import Data.Vec as Vec using (Vec; []; _∷_; head; tail) 10 | 11 | open import Relation.Binary using (Rel) 12 | open import Relation.Binary.PropositionalEquality 13 | open import Relation.Binary.Product.Pointwise 14 | 15 | open import NativeIO 16 | 17 | open import SizedIO.Base 18 | open import SizedIO.Console hiding (main) 19 | 20 | open import StateSizedIO.GUI.BaseStateDependent 21 | open import StateSizedIO.Object 22 | open import StateSizedIO.IOObject 23 | 24 | 25 | open import Size 26 | 27 | StackStateˢ = ℕ 28 | 29 | data StackMethodˢ (A : Set) : StackStateˢ → Set where 30 | push : {n : StackStateˢ} → A → StackMethodˢ A n 31 | pop : {n : StackStateˢ} → StackMethodˢ A (suc n) 32 | 33 | StackResultˢ : (A : Set) → (s : StackStateˢ) → StackMethodˢ A s → Set 34 | StackResultˢ A .n (push { n } x₁) = Unit 35 | StackResultˢ A (suc .n) (pop {n} ) = A 36 | 37 | nˢ : (A : Set) → (s : StackStateˢ) → (m : StackMethodˢ A s) → (r : StackResultˢ A s m) → StackStateˢ 38 | nˢ A .n (push { n } x) r = suc n 39 | nˢ A (suc .n) (pop { n }) r = n 40 | 41 | 42 | StackInterfaceˢ : (A : Set) → Interfaceˢ 43 | StackInterfaceˢ A .Stateˢ = StackStateˢ 44 | StackInterfaceˢ A .Methodˢ = StackMethodˢ A 45 | StackInterfaceˢ A .Resultˢ = StackResultˢ A 46 | StackInterfaceˢ A .nextˢ = nˢ A 47 | 48 | stackP : ∀{n : ℕ} → (i : Size) → (v : Vec String n) → IOObjectˢ consoleI (StackInterfaceˢ String) i n 49 | method (stackP { n } i es) {j} (push e) = return (_ , stackP j (e ∷ es)) 50 | method (stackP {suc n} i (x ∷ xs)){j} pop = return (x , stackP j xs) 51 | 52 | 53 | 54 | -- UNSIZED Version, without IO 55 | 56 | stackP' : ∀{n : ℕ} → (v : Vec String n) → Objectˢ (StackInterfaceˢ String) n 57 | stackP' es .objectMethod (push e) = (_ , stackP' (e ∷ es)) 58 | stackP' (x ∷ xs) .objectMethod pop = x , stackP' xs 59 | 60 | 61 | 62 | stackO : ∀{E : Set} {n : ℕ} (v : Vec E n) → Objectˢ (StackInterfaceˢ E) n 63 | objectMethod (stackO es) (push e) = _ , stackO (e ∷ es) 64 | objectMethod (stackO (e ∷ es)) pop = e , stackO es 65 | 66 | stackF : ∀{E : Set} (n : ℕ) (f : ℕ → E) → Objectˢ (StackInterfaceˢ E) n 67 | objectMethod (stackF n f) (push x) = _ , stackF (suc n) 68 | \{ zero → x 69 | ; (suc m) → f m 70 | } 71 | objectMethod (stackF (suc n) f) pop = (f zero) , stackF n (f ∘ suc) 72 | 73 | tabulate : ∀{E : Set} (n : ℕ) (f : ℕ → E) → Vec E n 74 | tabulate zero f = [] 75 | tabulate (suc n) f = f zero ∷ tabulate n λ m → f (suc m) 76 | 77 | module _ {E : Set} where 78 | private 79 | I = StackInterfaceˢ E 80 | S = Stateˢ I 81 | O = Objectˢ I 82 | open Bisim I 83 | 84 | _≡×≅'_ : ∀{s} → (o o' : E × O s) → Set 85 | _≡×≅'_ = _≡_ ×-Rel _≅_ 86 | 87 | Eq×Bisim : ∀ (s : S) → (o o' : E × O s) → Set 88 | Eq×Bisim s = _≡_ ×-Rel _≅_ 89 | 90 | pop-after-push : ∀{n}{v : Vec E n} (e : E) (let stack = stackO v) → 91 | (objectMethod stack (push e) ▹ λ { (_ , stack₁) → 92 | objectMethod stack₁ pop ▹ λ { (e₁ , stack₂) → 93 | ( e₁ , stack₂ ) }}) 94 | ≡×≅' (e , stack) 95 | pop-after-push e = refl , refl≅ _ 96 | 97 | 98 | push-after-pop : ∀{n}{v : Vec E n} (e : E) (let stack = stackO (e ∷ v)) → 99 | (objectMethod stack pop ▹ λ { (e₁ , stack₁) → 100 | objectMethod stack₁ (push e₁) ▹ λ { (_ , stack₂) → 101 | stack₂ }}) 102 | ≅ stack 103 | push-after-pop e = refl≅ _ 104 | 105 | 106 | -- The implementations of stacks with either vectors or functions are bisimilar. 107 | 108 | impl-bisim : ∀{n : ℕ} {f : ℕ → E} (v : Vec E n) (p : tabulate n f ≡ v) 109 | → stackF n f ≅ stackO v 110 | 111 | bisimMethod (impl-bisim v p) (push e) = 112 | bisim (impl-bisim (e ∷ v) (cong (_∷_ e) p)) 113 | 114 | bisimMethod (impl-bisim (x ∷ v) p) pop rewrite cong head p = 115 | bisim (impl-bisim v (cong tail p)) 116 | 117 | program : IOConsole ∞ Unit 118 | program = 119 | exec getLine λ str₀ → 120 | method s₀ (push str₀) >>= λ{ (_ , s₁) → -- empty 121 | exec getLine λ str₁ → 122 | method s₁ (push str₁) >>= λ{ (_ , s₂) → -- full 123 | method s₂ pop >>= λ{ (str₂ , s₃) → 124 | exec (putStrLn ("first pop: " Str.++ str₂) ) λ _ → 125 | exec getLine λ str₃ → 126 | method s₃ (push str₃) >>= λ{ (_ , s₄) → 127 | method s₄ pop >>= λ{ (str₄ , s₅) → 128 | exec (putStrLn ("second pop: " Str.++ str₄) ) λ _ → 129 | method s₅ pop >>= λ{ (str₅ , s₅) → 130 | exec (putStrLn ("third pop: " Str.++ str₅) ) λ _ → 131 | return unit 132 | }}}}}} 133 | where 134 | s₀ = stackP ∞ [] 135 | 136 | main : NativeIO Unit 137 | main = translateIOConsole program 138 | 139 | -------------------------------------------------------------------------------- /examples/StateSized/cellStateDependent.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --postfix-projections #-} 2 | 3 | module StateSized.cellStateDependent where 4 | 5 | open import Data.Product 6 | open import Data.String.Base 7 | 8 | {- 9 | open import SizedIO.Object 10 | open import SizedIO.ConsoleObject 11 | -} 12 | 13 | open import SizedIO.Console hiding (main) 14 | 15 | open import SizedIO.Base 16 | 17 | open import NativeIO 18 | 19 | open import StateSizedIO.Object 20 | open import StateSizedIO.IOObject 21 | 22 | 23 | open import Size 24 | open import StateSizedIO.cellStateDependent 25 | 26 | 27 | {- I moved most into src/StateSizedIO/cellStateDependent.agda 28 | 29 | now the code doesn't work 30 | -} 31 | 32 | -- Program is another program 33 | program : IOConsole ∞ Unit 34 | program = 35 | let c₀ = cellPempty ∞ in 36 | exec getLine λ str → 37 | method c₀ (put str) >>= λ{ (_ , c₁) → -- empty 38 | method c₁ get >>= λ{ (str₁ , c₂) → -- full 39 | exec (putStrLn ("we got " ++ str₁)) λ _ → 40 | exec (putStrLn ("Second Round")) λ _ → 41 | exec getLine λ str₂ → 42 | method c₂ (put str₂) >>= λ{ (_ , c₃) → 43 | method c₃ get >>= λ{ (str₃ , c₄) → 44 | exec (putStrLn ("we got " ++ str₃) ) λ _ → 45 | return unit 46 | }}}} 47 | 48 | main : NativeIO Unit 49 | main = translateIOConsole program 50 | 51 | 52 | 53 | -- cellP is constructor for the consoleObject for interface (cellI String) 54 | {- 55 | cellPˢ : ∀{i} → CellCˢ i 56 | force (method cellPˢ (put x)) = {!!} 57 | -} 58 | 59 | {- 60 | cellP : ∀{i} (s : String) → CellC i 61 | force (method (cellP s) get) = 62 | exec' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 63 | return (s , cellP s) 64 | force (method (cellP s) (put x)) = 65 | exec' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 66 | return (_ , (cellP x)) 67 | -} 68 | 69 | 70 | {- 71 | -- cellI is the interface of the object of a simple cell 72 | cellI : (A : Set) → Interfaceˢ 73 | Method (cellI A) = CellMethodˢ A 74 | Result (cellI A) m = CellResultˢ m 75 | 76 | -- cellC is the type of consoleObjects with interface (cellI String) 77 | CellC : (i : Size) → Set 78 | CellC i = ConsoleObject i (cellI String) 79 | -} 80 | {- 81 | -- cellO is a program for a simple cell which 82 | -- when get is called writes "getting s" for the string s of the object 83 | -- and when putting s writes "putting s" for the string 84 | 85 | -- cellP is constructor for the consoleObject for interface (cellI String) 86 | cellP : ∀{i} (s : String) → CellC i 87 | force (method (cellP s) get) = 88 | exec' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 89 | return (s , cellP s) 90 | force (method (cellP s) (put x)) = 91 | exec' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 92 | return (_ , (cellP x)) 93 | 94 | -- Program is another program 95 | program : String → IOConsole ∞ Unit 96 | program arg = 97 | let c₀ = cellP "Start" in 98 | method c₀ get >>= λ{ (s , c₁) → 99 | exec1 (putStrLn s) >> 100 | method c₁ (put arg) >>= λ{ (_ , c₂) → 101 | method c₂ get >>= λ{ (s' , c₃) → 102 | exec1 (putStrLn s') }}} 103 | 104 | main : NativeIO Unit 105 | main = translateIOConsole (program "hello") 106 | -} 107 | -------------------------------------------------------------------------------- /examples/UnSized/Console.agda: -------------------------------------------------------------------------------- 1 | module UnSized.Console where 2 | 3 | open import UnSizedIO.Base hiding (main) 4 | open import UnSizedIO.Console hiding (main) 5 | 6 | open import NativeIO 7 | 8 | {-# TERMINATING #-} 9 | myProgram : IOConsole Unit 10 | force myProgram = exec' getLine λ line → 11 | delay (exec' (putStrLn line) λ _ → 12 | delay (exec' (putStrLn line) λ _ → 13 | myProgram 14 | )) 15 | 16 | 17 | main : NativeIO Unit 18 | main = translateIOConsole myProgram 19 | -------------------------------------------------------------------------------- /examples/UnSized/Parrot.agda: -------------------------------------------------------------------------------- 1 | module UnSized.Parrot where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import UnSizedIO.IOObject 7 | 8 | open import UnSizedIO.Base hiding (main) 9 | open import UnSizedIO.Console hiding (main) 10 | 11 | open import NativeIO 12 | 13 | open import UnSized.SimpleCell hiding (program; main) 14 | 15 | 16 | --ParrotC : Set 17 | --ParrotC = ConsoleObject (cellI String) 18 | 19 | -- but reusing cellC from SimpleCell, as interface is ident! 20 | 21 | -- class Parrot implements Cell { 22 | -- Cell cell; 23 | -- Parrot (Cell c) { cell = c; } 24 | -- public String get() { 25 | -- return "(" + cell.get() + ") is what parrot got"; 26 | -- } 27 | -- public void put (String s) { 28 | -- cell.put("parrot puts (" + s + ")"); 29 | -- } 30 | -- } 31 | 32 | -- parrotP is constructor for the consoleObject for interface (cellI String) 33 | {-# NON_TERMINATING #-} 34 | parrotP : (c : CellC) → CellC 35 | (method (parrotP c) get) = 36 | exec1 (putStrLn "printing works") >> 37 | method c get >>= λ { (s , c') → 38 | return ("(" ++ s ++ ") is what parrot got" , parrotP c') } 39 | (method (parrotP c) (put s)) = 40 | method c (put ("parrot puts (" ++ s ++ ")")) 41 | 42 | 43 | 44 | -- public static void main (String[] args) { 45 | -- Parrot c = new Parrot(new SimpleCell("Start")); 46 | -- System.out.println(c.get()); 47 | -- c.put(args[1]); 48 | -- System.out.println(c.get()); 49 | -- } 50 | -- } 51 | 52 | 53 | program : String → IOConsole Unit 54 | program arg = 55 | let c₀ = parrotP (cellP "Start") in 56 | method c₀ get >>= λ{ (s , c₁) → 57 | exec1 (putStrLn s) >> 58 | method c₁ (put arg) >>= λ{ (_ , c₂) → 59 | method c₂ get >>= λ{ (s' , c₃) → 60 | exec1 (putStrLn s') }}} 61 | 62 | main : NativeIO Unit 63 | main = translateIOConsole (program "hello") 64 | -------------------------------------------------------------------------------- /examples/UnSized/SelfRef.agda: -------------------------------------------------------------------------------- 1 | module UnSized.SelfRef where 2 | 3 | open import Data.Unit.Base 4 | open import Data.Product 5 | open import Data.String.Base 6 | open import Data.Sum using (_⊎_) renaming (inj₁ to inl; inj₂ to inr) 7 | 8 | open import Size 9 | --open import SimpleCell 10 | open import SizedIO.Object 11 | open import SizedIO.IOObject 12 | open import SizedIO.ConsoleObject 13 | -- open import PrimTypeHelpersSmall 14 | 15 | open import UnSizedIO.Base hiding (main) 16 | open import UnSizedIO.Console hiding (main) 17 | 18 | open import NativeIO 19 | --open import SimpleCell 20 | 21 | -- Object Alpha 22 | 23 | data AlphaMethod A : Set where 24 | print : AlphaMethod A 25 | set : A → AlphaMethod A 26 | m1 : AlphaMethod A 27 | m2 : AlphaMethod A 28 | 29 | AlphaResponses : {A : Set} (c : AlphaMethod A) → Set 30 | AlphaResponses _ = ⊤ 31 | 32 | 33 | alphaI : (A : Set) → Interface 34 | Method (alphaI A) = AlphaMethod A 35 | Result (alphaI A) m = AlphaResponses m 36 | 37 | alphaC : (i : Size) → Set 38 | alphaC i = ConsoleObject i (alphaI String) 39 | 40 | -- 41 | -- Self Referential: method 'm1' calls method 'm2' 42 | -- 43 | {- 44 | -- {-# NON_TERMINATING #-} 45 | alphaO : ∀{i} (s : String) → alphaC i 46 | method (alphaO s) print = 47 | exec (putStrLn s) >> 48 | return (_ , alphaO s) 49 | 50 | method (alphaO s) (set x) = 51 | return (_ , alphaO x) 52 | 53 | -- force (method (alphaO s) m1) = exec (putStrLn s) λ _ → 54 | -- method (alphaO s) m2 >>= λ{ (_ , c₀) → 55 | -- return (_ , c₀) } 56 | method (alphaO s) m1 = 57 | exec1 (putStrLn s) >> 58 | method (alphaO s) m2 >>= λ{ (_ , c₀) → 59 | return (_ , c₀) } 60 | method (alphaO s) m2 = 61 | return (_ , alphaO (s ++ "->m2")) 62 | 63 | 64 | program : String → IOConsole ∞ Unit 65 | program arg = 66 | let c₀ = alphaO ("start̄\n====\n\n") in 67 | method c₀ m1 >>= λ{ (_ , c₁) → --- ===> m1 called, but m2 prints out text 68 | method c₁ print >>= λ{ (_ , c₂) → 69 | exec1 (putStrLn "\n\n====\nend") }} 70 | 71 | main : NativeIO Unit 72 | main = translateIOConsole (program "") 73 | -} 74 | -------------------------------------------------------------------------------- /examples/UnSized/SimpleCell.agda: -------------------------------------------------------------------------------- 1 | module UnSized.SimpleCell where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import UnSizedIO.Object 7 | open import UnSizedIO.IOObject 8 | open import UnSizedIO.ConsoleObject 9 | 10 | open import UnSizedIO.Base hiding (main) 11 | open import UnSizedIO.Console hiding (main) 12 | 13 | open import NativeIO 14 | 15 | data CellMethod A : Set where 16 | get : CellMethod A 17 | put : A → CellMethod A 18 | 19 | CellResult : ∀{A} → CellMethod A → Set 20 | CellResult {A} get = A 21 | CellResult (put _) = Unit 22 | 23 | -- cellI is the interface of the object of a simple cell 24 | cellI : (A : Set) → Interface 25 | Method (cellI A) = CellMethod A 26 | Result (cellI A) m = CellResult m 27 | 28 | -- cellC is the type of consoleObjects with interface (cellI String) 29 | CellC : Set 30 | CellC = ConsoleObject (cellI String) 31 | 32 | 33 | -- cellO is a program for a simple cell which 34 | -- when get is called writes "getting s" for the string s of the object 35 | -- and when putting s writes "putting s" for the string 36 | 37 | -- cellP is constructor for the consoleObject for interface (cellI String) 38 | cellP : (s : String) → CellC 39 | force (method (cellP s) get) = 40 | exec' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 41 | delay (return' (s , cellP s)) 42 | force (method (cellP s) (put x)) = 43 | exec' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 44 | delay (return' (_ , (cellP x))) 45 | 46 | -- Program is another program 47 | 48 | program : String → IOConsole Unit 49 | program arg = 50 | let c₀ = cellP "Start" in 51 | method c₀ get >>= λ{ (s , c₁) → 52 | exec1 (putStrLn s) >> 53 | method c₁ (put arg) >>= λ{ (_ , c₂) → 54 | method c₂ get >>= λ{ (s' , c₃) → 55 | exec1 (putStrLn s') }}} 56 | 57 | main : NativeIO Unit 58 | main = translateIOConsole (program "hello") 59 | -------------------------------------------------------------------------------- /examples/consoleExamples/helloWorld.agda: -------------------------------------------------------------------------------- 1 | module consoleExamples.helloWorld where 2 | 3 | open import ConsoleLib 4 | 5 | main : ConsoleProg 6 | main = run (WriteString "Hello World") 7 | -------------------------------------------------------------------------------- /examples/consoleExamples/passwordCheckSimple.agda: -------------------------------------------------------------------------------- 1 | module consoleExamples.passwordCheckSimple where 2 | 3 | open import ConsoleLib 4 | open import Data.Bool.Base 5 | open import Data.Bool 6 | open import Data.String renaming (_==_ to _==str_) 7 | open import SizedIO.Base 8 | 9 | main : ConsoleProg 10 | main = run (GetLine >>= λ s → 11 | if s ==str "passwd" 12 | then WriteString "Success" 13 | else WriteString "Error") 14 | -------------------------------------------------------------------------------- /examples/consoleExamples/simpleLoop.agda: -------------------------------------------------------------------------------- 1 | module consoleExamples.simpleLoop where 2 | 3 | open import ConsoleLib 4 | open import Data.Bool renaming (not to ¬) 5 | open import Data.String hiding (_==_) 6 | open import SizedIO.Base 7 | open import Size 8 | 9 | -- the following program asks the user for an input, which is a string. 10 | -- It converts this string into a Boolean value (by mapping "true" to true and 11 | -- "false" to false), applies Boolean negation to the result 12 | -- and then converts the resulting Boolean value back to a string 13 | 14 | boolStrInput2Bool : String → Bool 15 | boolStrInput2Bool s = s ==str "true" 16 | 17 | checkStrBoolInput : String → Bool 18 | checkStrBoolInput s = (s ==str "true") ∨ (s ==str "false") 19 | 20 | 21 | bool2Str : Bool → String 22 | bool2Str true = "true" 23 | bool2Str false = "false" 24 | 25 | 26 | -- when writing a recursive program you need to write first the body 27 | -- as an element of IOConsole 28 | -- and apply an eliminator to it, followed by a a function which 29 | -- results in an element of IOConsole' 30 | -- otherwise the recursion would unfold and you get an infinite term 31 | -- 32 | -- The functions WriteString+ and ReadLine+ which need the continuing program 33 | -- as argument can be used to create elements of IOConsole' 34 | -- 35 | -- Furthermore after applying a size argument (i which is of type Size) 36 | -- the termination checker can figure out that your definition is productive 37 | -- i.e. has at least one interaction before carrying out the recursion. 38 | 39 | mainBody : ∀{i} → IOConsole i Unit 40 | force (mainBody) = WriteString+ "Enter true or false" 41 | (GetLine >>= λ s → 42 | if (checkStrBoolInput s) 43 | then (WriteString ("Result of negating your input is " ++ 44 | bool2Str (¬ (boolStrInput2Bool s))) >>= λ _ → 45 | mainBody) 46 | else 47 | (WriteString ("Please enter \"true\" or \"false\"") >>= λ _ → 48 | mainBody)) 49 | 50 | main : ConsoleProg 51 | main = run mainBody 52 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/BasicIO.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.BasicIO where 2 | 3 | open import Data.Maybe hiding ( _>>=_ ) 4 | open import Data.Sum renaming (inj₁ to left; inj₂ to right; [_,_]′ to either) 5 | open import Function 6 | 7 | open import examplesPaperJFP.NativeIOSafe 8 | open import examplesPaperJFP.triangleRightOperator 9 | 10 | module _ where -- for indentation 11 | record IOInterface : Set₁ where 12 | field Command : Set 13 | Response : (c : Command) → Set 14 | 15 | 16 | open IOInterface public 17 | 18 | module _ {C : Set} {R : C → Set} (let 19 | I = record { Command = C; Response = R } 20 | ) where 21 | module Postulated {I : IOInterface} (let C = Command I) (let R = Response I) where 22 | 23 | postulate 24 | IO : (I : IOInterface) (A : Set) → Set 25 | exec : ∀{A} (c : C) (f : R c → IO I A) → IO I A 26 | return : ∀{A} (a : A) → IO I A 27 | _>>=_ : ∀{A B} (m : IO I A) (k : A → IO I B) → IO I B 28 | 29 | mutual 30 | record IO I A : Set where 31 | coinductive 32 | constructor delay 33 | field force : IO′ I A 34 | 35 | data IO′ I A : Set where 36 | exec′ : (c : Command I) (f : Response I c → IO I A) → IO′ I A 37 | return′ : (a : A) → IO′ I A 38 | 39 | open IO public 40 | 41 | delay′ : ∀{I A} → IO′ I A → IO I A 42 | force (delay′ x) = x 43 | 44 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) where 45 | -- module _ {I : IOInterface} (let record { Command = C; Response = R } = I) where 46 | -- module _ {I : IOInterface} (let ioInterface C R = I) where 47 | exec : ∀{A} (c : C) (f : R c → IO I A) → IO I A 48 | force (exec c f) = exec′ c f 49 | 50 | return : ∀{A} (a : A) → IO I A 51 | force (return a) = return′ a 52 | 53 | infixl 2 _>>=_ 54 | _>>=_ : ∀{A B} (m : IO I A) (k : A → IO I B) → IO I B 55 | force (m >>= k) with force m 56 | ... | exec′ c f = exec′ c λ x → f x >>= k 57 | ... | return′ a = force (k a) 58 | 59 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) 60 | where 61 | 62 | 63 | 64 | 65 | {-# NON_TERMINATING #-} 66 | translateIO : ∀ {A} (tr : (c : C) → NativeIO (R c)) → IO I A → NativeIO A 67 | translateIO tr m = force m ▹ λ 68 | { (exec′ c f) → (tr c) native>>= λ r → translateIO tr (f r) 69 | ; (return′ a) → nativeReturn a 70 | } 71 | 72 | -- Recursion 73 | 74 | -- trampoline provides a generic form of loop (generalizing while/repeat). 75 | -- Starting at state s : S, step function f is iterated until it returns 76 | -- a result in A. 77 | 78 | module _ (I : IOInterface)(let C = Command I) (let R = Response I) 79 | where 80 | data IO+ (A : Set) : Set where 81 | exec′ : (c : C) (f : R c → IO I A) → IO+ A 82 | 83 | module _ {I : IOInterface}(let C = Command I) (let R = Response I) 84 | where 85 | fromIO+′ : ∀{A} → IO+ I A → IO′ I A 86 | fromIO+′ (exec′ c f) = exec′ c f 87 | 88 | fromIO+ : ∀{A} → IO+ I A → IO I A 89 | force (fromIO+ (exec′ c f)) = exec′ c f 90 | 91 | _>>=+′_ : ∀{A B} (m : IO+ I A) (k : A → IO I B) → IO′ I B 92 | exec′ c f >>=+′ k = exec′ c λ x → f x >>= k 93 | 94 | _>>=+_ : ∀{A B} (m : IO+ I A) (k : A → IO I B) → IO I B 95 | force (m >>=+ k) = m >>=+′ k 96 | 97 | mutual 98 | 99 | _>>+_ : ∀{A B} (m : IO I (A ⊎ B)) (k : A → IO I B) → IO I B 100 | force (m >>+ k) = force m >>+′ k 101 | 102 | _>>+′_ : ∀{A B} (m : IO′ I (A ⊎ B)) (k : A → IO I B) → IO′ I B 103 | exec′ c f >>+′ k = exec′ c λ x → f x >>+ k 104 | return′ (left a) >>+′ k = force (k a) 105 | return′ (right b) >>+′ k = return′ b 106 | 107 | -- loop 108 | 109 | {-# TERMINATING #-} 110 | 111 | trampoline : ∀{A S} (f : S → IO+ I (S ⊎ A)) (s : S) → IO I A 112 | force (trampoline f s) = case (f s) of 113 | \ { (exec′ c k) → exec′ c λ r → k r >>+ trampoline f } 114 | 115 | -- simple infinite loop 116 | 117 | {-# TERMINATING #-} 118 | forever : ∀{A B} → IO+ I A → IO I B 119 | force (forever (exec′ c f)) = exec′ c λ r → f r >>= λ _ → forever (exec′ c f) 120 | 121 | whenJust : {A : Set} → Maybe A → (A → IO I Unit) → IO I Unit 122 | whenJust nothing k = return _ 123 | whenJust (just a) k = k a 124 | 125 | 126 | main : NativeIO Unit 127 | main = nativePutStrLn "Hello, world!" 128 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/CatNonTerm.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.CatNonTerm where 2 | 3 | open import examplesPaperJFP.BasicIO hiding (main) 4 | open import examplesPaperJFP.Console hiding (main) 5 | open import examplesPaperJFP.NativeIOSafe 6 | module _ where 7 | {-# NON_TERMINATING #-} 8 | cat : IO ConsoleInterface Unit 9 | cat = exec getLine λ{ nothing → return unit ; 10 | (just line) → 11 | exec (putStrLn line) λ _ → 12 | cat } 13 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/CatTerm.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.CatTerm where 2 | 3 | open import examplesPaperJFP.BasicIO hiding (main) 4 | open import examplesPaperJFP.Console hiding (main) 5 | open import examplesPaperJFP.NativeIOSafe 6 | 7 | cat : IO ConsoleInterface Unit 8 | force cat = 9 | exec′ getLine λ{ nothing → return unit ; (just line) → delay ( 10 | exec′ (putStrLn line) λ _ → 11 | cat )} 12 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Coalgebra.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Coalgebra where 2 | 3 | open import Size 4 | 5 | F : Set → Set 6 | mapF : ∀{A B} (f : A → B) → (F A → F B) 7 | 8 | --- Dummy implementation to satisfy Agda's positivity checker. 9 | F X = X 10 | mapF f x = f x 11 | 12 | S : Set 13 | t : S → F S 14 | 15 | data S′ : Set where 16 | 17 | S = S′ 18 | t x = x 19 | 20 | record νF : Set where 21 | coinductive 22 | field force : F νF 23 | 24 | open νF using (force) 25 | 26 | {-# TERMINATING #-} 27 | unfoldF : ∀{S} (t : S → F S) → (S → νF) 28 | force (unfoldF t s) = mapF (unfoldF t) (t s) 29 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Colists.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Colists where 2 | 3 | -- open import Function 4 | 5 | -- Case expressions (to be used with pattern-matching lambdas, see 6 | -- README.Case). 7 | 8 | infix 0 case_return_of_ case_of_ 9 | 10 | case_return_of_ : 11 | ∀ {a b} {A : Set a} 12 | (x : A) (B : A → Set b) → ((x : A) → B x) → B x 13 | case x return B of f = f x 14 | 15 | case_of_ : ∀ {a b} {A : Set a} {B : Set b} → A → (A → B) → B 16 | case x of f = case x return _ of f 17 | 18 | data ListF A S : Set where 19 | nil : ListF A S 20 | cons : (a : A) (s : S) → ListF A S 21 | 22 | record Colist A : Set where 23 | coinductive 24 | field force : ListF A (Colist A) 25 | 26 | 27 | open Colist using (force) 28 | 29 | 30 | mapF : ∀{A S S′} (f : S → S′) → (ListF A S → ListF A S′) 31 | mapF f nil = nil 32 | mapF f (cons a s) = cons a (f s) 33 | 34 | -- As an example, we define mapping over potentially infinite lists using 35 | -- copattern matching: 36 | 37 | map : ∀{A B} (f : A → B) (l : Colist A) → Colist B 38 | force (map f l) with force l 39 | ... | nil = nil 40 | ... | cons x xs = cons (f x) (map f xs) 41 | 42 | unfold : ∀{A S} (t : S → ListF A S) → (S → Colist A) 43 | force (unfold t s) with t s 44 | ... | nil = nil 45 | ... | cons a s′ = cons a (unfold t s′) 46 | 47 | fmap : ∀{A B} (f : A → B) (l : Colist A) → Colist B 48 | fmap f = unfold λ l → case force l of λ 49 | { nil → nil 50 | ; (cons a s) → cons (f a) s 51 | } 52 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Collatz.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Collatz where 2 | 3 | open import Data.Nat.Base 4 | open import Data.Nat.DivMod 5 | open import Data.Fin using (Fin; zero; suc) 6 | 7 | open import examplesPaperJFP.Colists 8 | 9 | collatzStep : ℕ → ListF ℕ ℕ 10 | collatzStep 1 = nil 11 | collatzStep n with n divMod 2 12 | ... | result q zero _ = cons n q 13 | ... | _ = cons n (1 + 3 * n) 14 | 15 | collatzSequence : ℕ → Colist ℕ 16 | collatzSequence = unfold collatzStep 17 | 18 | open Colist 19 | open import Data.List 20 | 21 | displayList : Colist ℕ → ℕ → List ℕ 22 | displayList s 0 = [] 23 | displayList s (suc m) with force s 24 | ... | nil = [] 25 | ... | (cons k s′) = k ∷ displayList s′ m 26 | 27 | 28 | displayCollatz : ℕ → ℕ → List ℕ 29 | displayCollatz n m = displayList (collatzSequence n) m 30 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Console.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Console where 2 | 3 | open import examplesPaperJFP.NativeIOSafe 4 | open import examplesPaperJFP.BasicIO hiding (main) 5 | open import examplesPaperJFP.ConsoleInterface public 6 | 7 | IOConsole : Set → Set 8 | IOConsole = IO ConsoleInterface 9 | 10 | --IOConsole+ : Set → Set 11 | --IOConsole+ = IO+ ConsoleInterface 12 | 13 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 14 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 15 | translateIOConsoleLocal getLine = nativeGetLine 16 | 17 | translateIOConsole : {A : Set} → IOConsole A → NativeIO A 18 | translateIOConsole = translateIO translateIOConsoleLocal 19 | 20 | main : NativeIO Unit 21 | main = nativePutStrLn "Console" 22 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/ConsoleInterface.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.ConsoleInterface where 2 | 3 | open import examplesPaperJFP.NativeIOSafe 4 | open import examplesPaperJFP.BasicIO hiding (main) 5 | 6 | module _ where 7 | data ConsoleCommand : Set where 8 | getLine : ConsoleCommand 9 | putStrLn : String → ConsoleCommand 10 | 11 | ConsoleResponse : ConsoleCommand → Set 12 | ConsoleResponse getLine = Maybe String 13 | ConsoleResponse (putStrLn s) = Unit 14 | 15 | ConsoleInterface : IOInterface 16 | Command ConsoleInterface = ConsoleCommand 17 | Response ConsoleInterface = ConsoleResponse 18 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/CounterCell.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.CounterCell where 2 | 3 | open import Data.Product 4 | open import Data.Nat.Base 5 | open import Data.Nat.Show 6 | open import Data.String.Base using (String; _++_) 7 | 8 | open import Size 9 | 10 | -- open import NativeIO 11 | 12 | -- open import SizedIO.Base 13 | -- open import SizedIO.Object 14 | -- open import SizedIO.IOObject 15 | -- open import SizedIO.ConsoleObject 16 | -- open import SizedIO.Console hiding (main) 17 | 18 | open import examplesPaperJFP.NativeIOSafe 19 | 20 | open import examplesPaperJFP.ConsoleInterface 21 | open import examplesPaperJFP.Console hiding (main) 22 | open import examplesPaperJFP.Object using (Interface; Method; Result; CellMethod; get; put; CellResult; cellJ) 23 | open import examplesPaperJFP.Sized hiding (program; main) 24 | 25 | data CounterMethod A : Set where 26 | super : (m : CellMethod A) → CounterMethod A 27 | stats : CounterMethod A 28 | 29 | pattern getᶜ = super get 30 | pattern putᶜ x = super (put x) 31 | 32 | statsCellI : (A : Set) → Interface 33 | Method (statsCellI A) = CounterMethod A 34 | Result (statsCellI A) (super m) = Result (cellJ A) m 35 | Result (statsCellI A) stats = Unit 36 | 37 | CounterC : (i : Size) → Set 38 | CounterC = IOObject ConsoleInterface (statsCellI String) 39 | 40 | counterCell : ∀{i} (c : CellC i) (ngets nputs : ℕ) → CounterC i 41 | 42 | method (counterCell c ngets nputs) getᶜ = 43 | method c get >>= λ { (s , c′) → 44 | return (s , counterCell c′ (1 + ngets) nputs) } 45 | 46 | method (counterCell c ngets nputs) (putᶜ x) = 47 | method c (put x) >>= λ { (_ , c′) → 48 | return (unit , counterCell c′ ngets (1 + nputs)) } 49 | 50 | method (counterCell c ngets nputs) stats = 51 | exec (putStrLn ("Counted " 52 | ++ show ngets ++ " calls to get and " 53 | ++ show nputs ++ " calls to put.")) λ _ → 54 | return (unit , counterCell c ngets nputs) 55 | 56 | program : String → IO ConsoleInterface ∞ Unit 57 | program arg = 58 | let c₀ = counterCell (simpleCell "Start") 0 0 in 59 | method c₀ getᶜ >>= λ{ (s , c₁) → 60 | exec (putStrLn s) λ _ → 61 | method c₁ (putᶜ arg) >>= λ{ (_ , c₂) → 62 | method c₂ getᶜ >>= λ{ (s′ , c₃) → 63 | exec (putStrLn s′) λ _ → 64 | method c₃ (putᶜ "Over!") >>= λ{ (_ , c₄) → 65 | method c₄ stats >>= λ{ (_ , c₅) → 66 | return unit }}}}} 67 | 68 | main : NativeIO Unit 69 | main = translateIO translateIOConsoleLocal (program "Hello") 70 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Equality.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Equality where 2 | 3 | infix 4 _≡_ 4 | 5 | data _≡_ {a} {A : Set a} (x : A) : A → Set a where 6 | refl : x ≡ x 7 | 8 | -- obsolete, now in Agda.Builtin.Equality: {-# BUILTIN EQUALITY _≡_ #-} 9 | -- No longer exists in Agda: {-# BUILTIN REFL refl #-} 10 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/ExampleDrawingProgram.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module examplesPaperJFP.ExampleDrawingProgram where 4 | 5 | open import Size 6 | open import Data.Bool.Base 7 | open import Data.Char.Base 8 | open import Data.Nat.Base hiding (_⊓_; _+_; _*_) 9 | open import Data.List.Base hiding (_++_) 10 | open import Data.Integer.Base hiding (suc) 11 | open import Data.String.Base 12 | open import Data.Maybe.Base 13 | open import Agda.Builtin.Char using (primCharEquality) 14 | 15 | open import Function 16 | 17 | open import SizedIO.Base 18 | open import SizedIO.IOGraphicsLib hiding (translateIOGraphics) renaming (translateIOGraphicsLocal to translateNative) 19 | 20 | open import NativeInt --PrimTypeHelpers 21 | open import NativeIO 22 | 23 | open import examplesPaperJFP.triangleRightOperator 24 | 25 | integerLess : ℤ → ℤ → Bool 26 | integerLess x y with ∣(y - (x ⊓ y))∣ 27 | ... | zero = true 28 | ... | z = false 29 | 30 | line : Point → Point → Graphic 31 | line p newpoint = withColor red (polygon 32 | (nativePoint x y 33 | ∷ nativePoint a b 34 | ∷ nativePoint (a + xAddition) (b + yAddition) 35 | ∷ nativePoint (x + xAddition) (y + yAddition) 36 | ∷ [] ) ) 37 | where 38 | x = nativeProj1Point p 39 | y = nativeProj2Point p 40 | a = nativeProj1Point newpoint 41 | b = nativeProj2Point newpoint 42 | 43 | diffx = + ∣ (a - x) ∣ 44 | diffy = + ∣ (b - y) ∣ 45 | 46 | diffx*3 = diffx * (+ 3) 47 | diffy*3 = diffy * (+ 3) 48 | 49 | condition = (integerLess diffx diffy*3) ∧ (integerLess diffy diffx*3) 50 | 51 | xAddition = if condition then + 2 else + 1 52 | yAddition = if condition then + 2 else + 1 53 | 54 | State = Maybe Point 55 | 56 | 57 | mutual 58 | 59 | loop : ∀{i} → Window → State → IOGraphics i Unit 60 | force (loop w s) = exec' (getWindowEvent w) λ e → 61 | winEvtHandler w s e 62 | 63 | winEvtHandler : ∀{i} → Window → State → Event → IOGraphics i Unit 64 | winEvtHandler w s (Key c t) = if primCharEquality c 'x' 65 | then (exec (closeWindow w) return) 66 | else loop w s 67 | winEvtHandler w s (MouseMove p₂) = s ▹ λ 68 | { nothing → loop w (just p₂) 69 | ; (just p₁) → exec (drawInWindow w (line p₁ p₂)) λ _ → 70 | loop w (just p₂) } 71 | winEvtHandler w s _ = loop w s 72 | 73 | 74 | program : ∀{i} → IOGraphics i Unit 75 | program = 76 | exec (openWindow "Drawing Prog" nothing 1000 1000 77 | nativeDrawGraphic nothing) λ win → 78 | loop win nothing 79 | 80 | translateIOGraphics : IOGraphics ∞ Unit → NativeIO Unit 81 | translateIOGraphics = translateIO translateNative 82 | 83 | main : NativeIO Unit 84 | main = nativeRunGraphics (translateIOGraphics program) 85 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/IOExampleConsole.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.IOExampleConsole where 2 | 3 | open import examplesPaperJFP.BasicIO hiding (main) 4 | open import examplesPaperJFP.Console hiding (main) 5 | open import examplesPaperJFP.NativeIOSafe 6 | open import examplesPaperJFP.CatTerm 7 | 8 | main : NativeIO Unit 9 | main = translateIOConsole cat 10 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/IOGraphicsLib.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module examplesPaperJFP.IOGraphicsLib where 4 | 5 | 6 | open import Data.Bool.Base 7 | open import Data.Char.Base 8 | open import Data.Nat.Base hiding (_⊓_; _+_; _*_) 9 | open import Data.List.Base hiding (_++_) 10 | open import Data.Integer.Base hiding (suc) 11 | open import Data.String.Base 12 | open import Data.Maybe.Base 13 | 14 | open import Function 15 | 16 | open import Size 17 | 18 | open import SizedIO.Base 19 | open import SizedIO.IOGraphicsLib hiding (GraphicsCommands; GraphicsResponses; GraphicsInterface; IOGraphics) renaming (Size to WindowSize) 20 | 21 | open import NativeInt --PrimTypeHelpers 22 | open import NativeIO 23 | 24 | 25 | 26 | data GraphicsCommands : Set where 27 | getWindowEvent : Window → GraphicsCommands 28 | openWindow : String → Maybe Point → ℕ → ℕ 29 | → RedrawMode → Maybe Word32 30 | → GraphicsCommands 31 | closeWindow : Window → GraphicsCommands 32 | drawInWindow : Window → Graphic → GraphicsCommands 33 | 34 | 35 | GraphicsResponses : GraphicsCommands → Set 36 | GraphicsResponses (getWindowEvent _) = Event 37 | GraphicsResponses (openWindow _ _ _ _ _ _) = Window 38 | GraphicsResponses (closeWindow _) = Unit 39 | GraphicsResponses (drawInWindow _ _) = Unit 40 | 41 | GraphicsInterface : IOInterface 42 | Command GraphicsInterface = GraphicsCommands 43 | Response GraphicsInterface = GraphicsResponses 44 | 45 | IOGraphics : Size → Set → Set 46 | IOGraphics i = IO GraphicsInterface i 47 | 48 | 49 | translateNative : (c : GraphicsCommands) → NativeIO (GraphicsResponses c) 50 | 51 | translateNative (getWindowEvent w) = nativeGetWindowEvent w 52 | translateNative (openWindow str point n1 n2 m w) = nativeOpenWindowExNat str point n1 n2 m w 53 | translateNative (closeWindow w) = nativeCloseWindow w 54 | translateNative (drawInWindow w g) = nativeDrawInWindow w g 55 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/NativeIOSafe.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.NativeIOSafe where 2 | 3 | open import Data.Maybe.Base using (Maybe; nothing; just) public 4 | open import Data.String.Base using (String) public 5 | 6 | record Unit : Set where 7 | constructor unit 8 | 9 | {-# COMPILE GHC Unit = () #-} 10 | 11 | postulate 12 | NativeIO : Set → Set 13 | nativeReturn : {A : Set} → A → NativeIO A 14 | _native>>=_ : {A B : Set} → NativeIO A → (A → NativeIO B) → NativeIO B 15 | 16 | 17 | {-# BUILTIN IO NativeIO #-} 18 | {-# COMPILE GHC NativeIO = IO #-} -- IO.FF-I.AgdaIO 19 | {-# COMPILE GHC _native>>=_ = (\_ _ -> (>>=) :: IO a -> (a -> IO b) -> IO b) #-} 20 | {-# COMPILE GHC nativeReturn = (\_ -> return :: a -> IO a) #-} 21 | 22 | postulate 23 | nativeGetLine : NativeIO (Maybe String) 24 | nativePutStrLn : String → NativeIO Unit 25 | 26 | {-# FOREIGN GHC import Data.Text #-} 27 | {-# FOREIGN GHC import System.IO.Error #-} 28 | 29 | {-# COMPILE GHC nativePutStrLn = (\ s -> putStrLn (Data.Text.unpack s)) #-} 30 | {-# COMPILE GHC nativeGetLine = (fmap (Just . Data.Text.pack) getLine `System.IO.Error.catchIOError` \ err -> if System.IO.Error.isEOFError err then return Nothing else ioError err) #-} 31 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Object.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Object where 2 | 3 | open import Data.Product 4 | open import Data.String.Base 5 | 6 | open import examplesPaperJFP.NativeIOSafe 7 | open import examplesPaperJFP.BasicIO hiding (main) 8 | open import examplesPaperJFP.Console hiding (main) 9 | 10 | record Interface : Set₁ where 11 | field Method : Set 12 | Result : (m : Method) → Set 13 | 14 | open Interface public 15 | 16 | record Object (I : Interface) : Set where 17 | coinductive 18 | field objectMethod : (m : Method I) → Result I m × Object I 19 | 20 | 21 | open Object public 22 | 23 | record IOObject (Iᵢₒ : IOInterface) (I : Interface) : Set where 24 | coinductive 25 | field method : (m : Method I) → IO Iᵢₒ (Result I m × IOObject Iᵢₒ I) 26 | 27 | open IOObject public 28 | 29 | 30 | 31 | data CellMethod A : Set where 32 | get : CellMethod A 33 | put : A → CellMethod A 34 | 35 | CellResult : ∀{A} → CellMethod A → Set 36 | CellResult {A} get = A 37 | CellResult (put _) = Unit 38 | 39 | cellJ : (A : Set) → Interface 40 | Method (cellJ A) = CellMethod A 41 | Result (cellJ A) m = CellResult m 42 | 43 | CellC : Set 44 | CellC = IOObject ConsoleInterface (cellJ String) 45 | 46 | simpleCell : (s : String) → CellC 47 | force (method (simpleCell s) get) = 48 | exec′ (putStrLn ("getting (" ++ s ++ ")")) λ _ → 49 | delay (return′ (s , simpleCell s)) 50 | force (method (simpleCell s) (put x)) = 51 | exec′ (putStrLn ("putting (" ++ x ++ ")")) λ _ → 52 | delay (return′ (unit , simpleCell x)) 53 | 54 | {-# TERMINATING #-} 55 | 56 | program : IOConsole Unit 57 | force program = 58 | let c₁ = simpleCell "Start" in 59 | exec′ getLine λ{ nothing → return unit; (just s) → 60 | method c₁ (put s) >>= λ{ (_ , c₂) → 61 | method c₂ get >>= λ{ (s′ , _ ) → 62 | exec (putStrLn s′) λ _ → 63 | program }}} 64 | 65 | main : NativeIO Unit 66 | main = translateIOConsole program 67 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/Sized.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.Sized where 2 | 3 | open import Data.Product using (_×_; _,_) 4 | open import Data.String 5 | 6 | open import Function using (case_of_) 7 | open import Size 8 | 9 | open import examplesPaperJFP.NativeIOSafe 10 | open import examplesPaperJFP.BasicIO using (IOInterface; Command; Response) 11 | open import examplesPaperJFP.ConsoleInterface 12 | open import examplesPaperJFP.Console using (translateIOConsoleLocal) 13 | open import examplesPaperJFP.Object using (Interface; Method; Result; 14 | cellJ; CellMethod; get; put; CellResult) 15 | 16 | 17 | module UnfoldF where 18 | open import examplesPaperJFP.Coalgebra using (F; mapF) 19 | record νF (i : Size) : Set where 20 | coinductive 21 | constructor delay 22 | field force : ∀(j : Size< i) → F (νF j) 23 | 24 | open νF using (force) 25 | 26 | unfoldF : ∀{S} (t : S → F S) → ∀ i → (S → νF i) 27 | force (unfoldF t i s) j = mapF (unfoldF t j) (t s) 28 | 29 | mutual 30 | record IO (Iᵢₒ : IOInterface) (i : Size) (A : Set) : Set where 31 | coinductive 32 | 33 | constructor delay 34 | field force : {j : Size< i} → IO′ Iᵢₒ j A 35 | 36 | data IO′ (Iᵢₒ : IOInterface) (i : Size) (A : Set) : Set where 37 | exec′ : (c : Command Iᵢₒ) (f : Response Iᵢₒ c → IO Iᵢₒ i A) → IO′ Iᵢₒ i A 38 | return′ : (a : A) → IO′ Iᵢₒ i A 39 | 40 | module NestedRecursion (Iᵢₒ : IOInterface) (A : Set) where 41 | 42 | data F (X : Set) : Set where 43 | exec′ : (c : Command Iᵢₒ) (f : Response Iᵢₒ c → X) → F X 44 | return′ : (a : A) → F X 45 | 46 | record νF (i : Size) : Set where 47 | coinductive 48 | constructor delay 49 | field force : {j : Size< i} → F (νF j) 50 | 51 | open IO public 52 | 53 | module _ {Iᵢₒ : IOInterface } (let C = Command Iᵢₒ) (let R = Response Iᵢₒ) where 54 | 55 | infixl 2 _>>=_ 56 | 57 | exec : ∀ {i A} (c : C) (f : R c → IO Iᵢₒ i A) → IO Iᵢₒ i A 58 | return : ∀ {i A} (a : A) → IO Iᵢₒ i A 59 | _>>=_ : ∀ {i A B} (m : IO Iᵢₒ i A) (k : A → IO Iᵢₒ i B) → IO Iᵢₒ i B 60 | 61 | force (exec c f) = exec′ c f 62 | force (return a) = return′ a 63 | 64 | force (_>>=_ {i} m k) {j} with force m {j} 65 | ... | exec′ c f = exec′ c λ r → _>>=_ {j} (f r) k 66 | ... | return′ a = force (k a) {j} 67 | 68 | 69 | {-# NON_TERMINATING #-} 70 | translateIO : ∀{A : Set} 71 | → (translateLocal : (c : C) → NativeIO (R c)) 72 | → IO Iᵢₒ ∞ A 73 | → NativeIO A 74 | translateIO translateLocal m = case (force m) of 75 | λ{ (exec′ c f) → (translateLocal c) native>>= λ r → 76 | translateIO translateLocal (f r) 77 | ; (return′ a) → nativeReturn a 78 | } 79 | 80 | record IOObject (Iᵢₒ : IOInterface) (I : Interface) (i : Size) : Set where 81 | coinductive 82 | field method : ∀{j : Size< i} (m : Method I) 83 | → IO Iᵢₒ ∞ (Result I m × IOObject Iᵢₒ I j) 84 | 85 | 86 | open IOObject public 87 | 88 | 89 | CellC : (i : Size) → Set 90 | CellC = IOObject ConsoleInterface (cellJ String) 91 | 92 | simpleCell : ∀{i} (s : String) → CellC i 93 | force (method (simpleCell {i} s) {j} get) = 94 | exec′ (putStrLn ("getting (" ++ s ++ ")")) λ _ → 95 | return (s , simpleCell {j} s) 96 | force (method (simpleCell _) (put s)) = 97 | exec′ (putStrLn ("putting (" ++ s ++ ")")) λ _ → 98 | return (unit , simpleCell s) 99 | 100 | program : ∀{i} → IO ConsoleInterface i Unit 101 | force program = 102 | let c₁ = simpleCell "Start" in 103 | exec′ getLine λ{ nothing → return unit; (just s) → 104 | method c₁ (put s) >>= λ{ (_ , c₂) → 105 | method c₂ get >>= λ{ (s′ , c₃) → 106 | exec (putStrLn s′) λ _ → 107 | program }}} 108 | 109 | 110 | main : NativeIO Unit 111 | main = translateIO translateIOConsoleLocal program 112 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/SpaceShipAdvanced.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.SpaceShipAdvanced where 2 | 3 | 4 | open import SizedIO.Base 5 | open import StateSizedIO.GUI.BaseStateDependent 6 | 7 | 8 | open import Data.Bool.Base 9 | open import Data.List.Base 10 | open import Data.Integer 11 | 12 | open import Data.Product hiding (map) 13 | open import SizedIO.Object 14 | open import SizedIO.IOObject 15 | 16 | open import NativeIO 17 | open import Size 18 | 19 | open import StateSizedIO.GUI.WxBindingsFFI 20 | open import StateSizedIO.GUI.VariableList 21 | 22 | open import StateSizedIO.GUI.WxGraphicsLib 23 | 24 | open import StateSized.GUI.BitMaps 25 | 26 | data GraphicServerMethod : Set where 27 | onPaintM : DC → Rect → GraphicServerMethod 28 | moveSpaceShipM : Frame → GraphicServerMethod 29 | callRepaintM : Frame → GraphicServerMethod 30 | 31 | GraphicServerResult : GraphicServerMethod → Set 32 | GraphicServerResult _ = Unit 33 | 34 | 35 | GraphicServerInterface : Interface 36 | Method GraphicServerInterface = GraphicServerMethod 37 | Result GraphicServerInterface = GraphicServerResult 38 | 39 | GraphicServerObject : ∀{i} → Set 40 | GraphicServerObject {i} = IOObject GuiLev1Interface GraphicServerInterface i 41 | 42 | graphicServerObject : ∀{i} → ℤ → GraphicServerObject {i} 43 | method (graphicServerObject z) (onPaintM dc rect) = 44 | exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 45 | return (unit , graphicServerObject z) 46 | method (graphicServerObject z) (moveSpaceShipM fra) = 47 | return (unit , graphicServerObject (z + (+ 20))) 48 | method (graphicServerObject z) (callRepaintM fra) = 49 | exec (repaint fra) λ _ → 50 | return (unit , graphicServerObject z) 51 | 52 | VarType = GraphicServerObject {∞} 53 | 54 | varInit : VarType 55 | varInit = graphicServerObject (+ 150) 56 | 57 | 58 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 59 | onPaint obj dc rect = mapIO proj₂ (method obj (onPaintM dc rect)) 60 | 61 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 62 | moveSpaceShip fra obj = mapIO proj₂ (method obj (moveSpaceShipM fra)) 63 | 64 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 65 | callRepaint fra obj = mapIO proj₂ (method obj (callRepaintM fra)) 66 | 67 | 68 | buttonHandler : ∀{i} → Frame → List (VarType → IO GuiLev1Interface i VarType) 69 | buttonHandler fra = moveSpaceShip fra ∷ [ callRepaint fra ] 70 | 71 | 72 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 73 | program = execˢ (level1C makeFrame) λ fra → 74 | execˢ (level1C (makeButton fra)) λ bt → 75 | execˢ (level1C (addButton fra bt)) λ _ → 76 | execˢ (createVar varInit) λ _ → 77 | execˢ (setButtonHandler bt 78 | (moveSpaceShip fra ∷ [ callRepaint fra ])) λ _ → 79 | execˢ (setOnPaint fra [ onPaint ]) 80 | returnˢ 81 | 82 | main : NativeIO Unit 83 | main = start (translateLev2 program) 84 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/SpaceShipCell.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.SpaceShipCell where 2 | 3 | 4 | open import SizedIO.Base 5 | open import StateSizedIO.GUI.BaseStateDependent 6 | 7 | 8 | open import Data.Bool.Base 9 | open import Data.List.Base 10 | open import Data.Integer 11 | 12 | open import Data.Product hiding (map) 13 | open import SizedIO.Object 14 | open import SizedIO.IOObject 15 | 16 | open import NativeIO 17 | open import Sized.SimpleCell hiding (main; program) 18 | 19 | open import StateSizedIO.GUI.WxBindingsFFI 20 | open import StateSizedIO.GUI.VariableList 21 | 22 | open import StateSizedIO.GUI.WxGraphicsLib 23 | 24 | open import StateSized.GUI.BitMaps 25 | 26 | 27 | VarType = Object (cellJ ℤ) 28 | 29 | cellℤC : (z : ℤ ) → VarType 30 | objectMethod (cellℤC z) get = ( z , cellℤC z ) 31 | objectMethod (cellℤC z) (put z′) = ( unit , cellℤC z′ ) 32 | 33 | varInit : VarType 34 | varInit = cellℤC (+ 150) 35 | 36 | 37 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 38 | 39 | 40 | onPaint c dc rect = 41 | let (z , c₁) = objectMethod c get in 42 | exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 43 | return c₁ 44 | 45 | moveSpaceShip : ∀{i} → Frame → VarType 46 | → IO GuiLev1Interface i VarType 47 | 48 | moveSpaceShip fra c = 49 | let (z , c₁) = objectMethod c get 50 | (_ , c₂) = objectMethod c₁ (put (z + (+ 20))) 51 | in return c₂ 52 | 53 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 54 | 55 | 56 | callRepaint fra c = exec (repaint fra) λ _ → return c 57 | 58 | 59 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 60 | program = execˢ (level1C makeFrame) λ fra → 61 | execˢ (level1C (makeButton fra)) λ bt → 62 | execˢ (level1C (addButton fra bt)) λ _ → 63 | execˢ (createVar varInit) λ _ → 64 | execˢ (setButtonHandler bt (moveSpaceShip fra 65 | ∷ [ callRepaint fra ])) λ _ → 66 | execˢ (setOnPaint fra [ onPaint ]) 67 | returnˢ 68 | 69 | main : NativeIO Unit 70 | main = start (translateLev2 program) 71 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/SpaceShipSimpleVar.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.SpaceShipSimpleVar where 2 | 3 | open import SizedIO.Base 4 | open import StateSizedIO.GUI.BaseStateDependent 5 | 6 | open import Data.Bool.Base 7 | open import Data.List.Base 8 | open import Data.Integer 9 | 10 | open import Data.Product hiding (map) 11 | open import SizedIO.Object 12 | open import SizedIO.IOObject 13 | 14 | open import NativeIO 15 | 16 | open import StateSizedIO.GUI.WxBindingsFFI 17 | open import StateSizedIO.GUI.VariableList 18 | 19 | open import StateSizedIO.GUI.WxGraphicsLib 20 | 21 | open import StateSized.GUI.BitMaps 22 | 23 | VarType = ℤ 24 | 25 | varInit : VarType 26 | varInit = (+ 150) 27 | 28 | 29 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 30 | onPaint z dc rect = exec (drawBitmap dc ship (z , (+ 150)) true) λ _ → 31 | return z 32 | 33 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 34 | moveSpaceShip fra z = return (z + (+ 20)) 35 | 36 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 37 | callRepaint fra z = exec (repaint fra) λ _ → return z 38 | 39 | 40 | buttonHandler : ∀{i} → Frame → List (VarType → IO GuiLev1Interface i VarType) 41 | buttonHandler fra = moveSpaceShip fra ∷ [ callRepaint fra ] 42 | 43 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 44 | program = execˢ (level1C makeFrame) λ fra → 45 | execˢ (level1C (makeButton fra)) λ bt → 46 | execˢ (level1C (addButton fra bt)) λ _ → 47 | execˢ (createVar varInit) λ _ → 48 | execˢ (setButtonHandler bt (moveSpaceShip fra 49 | ∷ [ callRepaint fra ])) λ _ → 50 | execˢ (setOnPaint fra [ onPaint ]) 51 | returnˢ 52 | 53 | main : NativeIO Unit 54 | main = start (translateLev2 program) 55 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/StackBisim.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.StackBisim where 2 | 3 | open import Data.Product 4 | open import Function 5 | open import Data.Nat.Base as N 6 | open import Data.Vec as Vec using (Vec; []; _∷_; head; tail) 7 | 8 | open import Relation.Binary.PropositionalEquality 9 | 10 | open import examplesPaperJFP.StatefulObject 11 | 12 | module _ {S : Set} {M : S → Set} {R : (s : S) → (m : M s) → Set} {next : (s : S) → (m : M s) → R s m → S} (let 13 | 14 | I = record { Stateˢ = S; Methodˢ = M; Resultˢ = R; nextˢ = next } 15 | ) (let 16 | O = Objectˢ I 17 | 18 | ) where 19 | module Bisim (I : Interfaceˢ) 20 | (let S = Stateˢ I) 21 | (let M = Methodˢ I) 22 | (let R = Resultˢ I) 23 | (let next = nextˢ I) 24 | (let O = Objectˢ I) where 25 | 26 | data ΣR {A : Set} {B : A → Set} (R : ∀{a} (b b′ : B a) → Set) 27 | : (p p′ : Σ[ a ∈ A ] B a) → Set 28 | where 29 | eqΣ : ∀{a}{b b′ : B a} → R b b′ → ΣR R (a , b) (a , b′) 30 | 31 | 32 | record _≅_ {s : S} (o o′ : O s) : Set where 33 | coinductive 34 | field bisimMethod : (m : M s) → 35 | ΣR (_≅_) (objectMethod o m) (objectMethod o′ m) 36 | 37 | open _≅_ public 38 | 39 | refl≅ : ∀{s} (o : O s) → o ≅ o 40 | bisimMethod (refl≅ o) m = let (r , o′) = objectMethod o m 41 | in eqΣ (refl≅ o′) 42 | 43 | module _ {E : Set} where 44 | private 45 | I = StackInterfaceˢ E 46 | S = Stateˢ I 47 | O = Objectˢ I 48 | open Bisim I 49 | 50 | pop-after-push : ∀{n} {v : Vec E n} {e : E} → 51 | let st = stack v 52 | (_ , st₁) = objectMethod st (push e) 53 | (e₂ , st₂) = objectMethod st₁ pop 54 | 55 | in (e ≡ e₂) × (st ≅ st₂) 56 | 57 | pop-after-push = refl , refl≅ _ 58 | 59 | 60 | push-after-pop : ∀{n} {v : Vec E n} {e : E} → 61 | let st = stack (e ∷ v) 62 | (e₁ , st₁) = objectMethod st pop 63 | (_ , st₂) = objectMethod st₁ (push e₁) 64 | 65 | in st ≅ st₂ 66 | 67 | push-after-pop = refl≅ _ 68 | 69 | 70 | tabulate : ∀ (n : ℕ) (f : ℕ → E) → Vec E n 71 | tabulate 0 f = [] 72 | tabulate (suc n) f = f 0 ∷ tabulate n (f ∘ suc) 73 | 74 | 75 | stackF : ∀ (n : ℕ) (f : ℕ → E) → Objectˢ (StackInterfaceˢ E) n 76 | objectMethod (stackF n f) (push e) = _ , stackF (suc n) λ 77 | { 0 → e 78 | ; (suc m) → f m } 79 | objectMethod (stackF (suc n) f) pop = f 0 , stackF n (f ∘ suc) 80 | 81 | 82 | impl-bisim : ∀{n f} v (p : tabulate n f ≡ v) → stackF n f ≅ stack v 83 | 84 | bisimMethod (impl-bisim v p) (push e) = 85 | eqΣ (impl-bisim (e ∷ v) (cong (_∷_ e) p)) 86 | 87 | bisimMethod (impl-bisim (e ∷ v) p) pop rewrite cong head p = 88 | eqΣ (impl-bisim v (cong tail p)) 89 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/StateDependentIO.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.StateDependentIO where 2 | 3 | open import Size 4 | open import NativeIO 5 | open import Function 6 | open import Agda.Primitive 7 | open import Level using (_⊔_) renaming (suc to lsuc) 8 | 9 | module _ {σ γ ρ} where 10 | 11 | record IOInterfaceˢ : Set (lsuc (σ ⊔ γ ⊔ ρ )) where 12 | field Stateˢ : Set σ 13 | Commandˢ : Stateˢ → Set γ 14 | Responseˢ : (s : Stateˢ) → Commandˢ s → Set ρ 15 | nextˢ : (s : Stateˢ) → (c : Commandˢ s) → Responseˢ s c → Stateˢ 16 | 17 | open IOInterfaceˢ public 18 | 19 | 20 | module _ {α σ γ ρ}(i : IOInterfaceˢ {σ} {γ} {ρ} ) 21 | (let S = Stateˢ i) (let C = Commandˢ i) 22 | (let R = Responseˢ i) (let next = nextˢ i) 23 | where 24 | mutual 25 | 26 | record IOˢ (i : Size) (A : S → Set α) (s : S) : Set (lsuc (α ⊔ σ ⊔ γ ⊔ ρ )) where 27 | coinductive 28 | 29 | constructor delay 30 | field forceˢ : {j : Size< i} → IOˢ′ j A s 31 | 32 | data IOˢ′ (i : Size) (A : S → Set α) : S → Set (lsuc (α ⊔ σ ⊔ γ ⊔ ρ )) where 33 | doˢ′ : {s : S} → (c : C s) → (f : (r : R s c) → IOˢ i A (next s c r) ) 34 | → IOˢ′ i A s 35 | returnˢ′ : {s : S} → (a : A s) → IOˢ′ i A s 36 | 37 | data IOˢ+ (i : Size) (A : S → Set α ) : S → Set (lsuc (α ⊔ σ ⊔ γ ⊔ ρ )) where 38 | doˢ′ : {s : S} → (c : C s) (f : (r : R s c) → IOˢ i A (next s c r)) 39 | → IOˢ+ i A s 40 | 41 | open IOˢ public 42 | 43 | module _ {α σ γ ρ}{I : IOInterfaceˢ {σ} {γ} {ρ}} 44 | (let S = Stateˢ I) (let C = Commandˢ I) 45 | (let R = Responseˢ I) (let next = nextˢ I) where 46 | 47 | returnˢ : ∀{i}{A : S → Set α} {s : S} (a : A s) → IOˢ I i A s 48 | forceˢ (returnˢ a) = returnˢ′ a 49 | 50 | doˢ : ∀{i}{A : S → Set α} {s : S} 51 | (c : C s) (f : (r : R s c) → IOˢ I i A (next s c r)) → IOˢ I i A s 52 | forceˢ (doˢ c f) = doˢ′ c f 53 | 54 | module _ {σ γ}{I : IOInterfaceˢ {σ} {γ} {lzero}} 55 | (let S = Stateˢ I) (let C = Commandˢ I) 56 | (let R = Responseˢ I) (let next = nextˢ I) where 57 | {-# NON_TERMINATING #-} 58 | 59 | translateIOˢ : ∀{A : Set }{s : S} 60 | → (translateLocal : (s : S) → (c : C s) → NativeIO (R s c)) 61 | → IOˢ I ∞ (λ s → A) s 62 | → NativeIO A 63 | 64 | translateIOˢ {A} {s} translateLocal p = case (forceˢ p {_}) of 65 | λ{ (doˢ′ {.s} c f) → (translateLocal s c) native>>= λ r → 66 | translateIOˢ translateLocal (f r) 67 | ; (returnˢ′ a) → nativeReturn a 68 | } 69 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/StatefulObject.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.StatefulObject where 2 | 3 | open import Data.Product 4 | open import Data.String.Base as Str 5 | open import Data.Nat.Base as N 6 | open import Data.Vec as Vec using (Vec; []; _∷_; head; tail) 7 | 8 | open import SizedIO.Console hiding (main) 9 | open import Size 10 | 11 | open import NativeIO 12 | 13 | open import SizedIO.Base 14 | 15 | 16 | StackStateˢ = ℕ 17 | 18 | record Interfaceˢ : Set₁ where 19 | field Stateˢ : Set 20 | Methodˢ : (s : Stateˢ) → Set 21 | Resultˢ : (s : Stateˢ) → (m : Methodˢ s) → Set 22 | nextˢ : (s : Stateˢ) → (m : Methodˢ s) → (r : Resultˢ s m) → Stateˢ 23 | 24 | open Interfaceˢ public 25 | 26 | data StackMethodˢ (A : Set) : (n : StackStateˢ) → Set where 27 | push : ∀ {n} → A → StackMethodˢ A n 28 | pop : ∀ {n} → StackMethodˢ A (suc n) 29 | 30 | StackResultˢ : (A : Set) → (s : StackStateˢ) → StackMethodˢ A s → Set 31 | StackResultˢ A _ (push _) = Unit 32 | StackResultˢ A _ pop = A 33 | 34 | 35 | stackNextˢ : ∀ A n (m : StackMethodˢ A n) (r : StackResultˢ A n m) → StackStateˢ 36 | stackNextˢ _ n (push _) _ = suc n 37 | stackNextˢ _ (suc n) pop _ = n 38 | 39 | StackInterfaceˢ : (A : Set) → Interfaceˢ 40 | Stateˢ (StackInterfaceˢ A) = StackStateˢ 41 | Methodˢ (StackInterfaceˢ A) = StackMethodˢ A 42 | Resultˢ (StackInterfaceˢ A) = StackResultˢ A 43 | nextˢ (StackInterfaceˢ A) = stackNextˢ A 44 | 45 | record Objectˢ (I : Interfaceˢ) (s : Stateˢ I) : Set where 46 | coinductive 47 | field objectMethod : (m : Methodˢ I s) → 48 | Σ[ r ∈ Resultˢ I s m ] Objectˢ I (nextˢ I s m r) 49 | open Objectˢ public 50 | 51 | record IOObjectˢ (Iᵢₒ : IOInterface) (I : Interfaceˢ) (s : Stateˢ I) : Set where 52 | coinductive 53 | field method : (m : Methodˢ I s) → 54 | IO Iᵢₒ ∞ (Σ[ r ∈ Resultˢ I s m ] IOObjectˢ Iᵢₒ I (nextˢ I s m r)) 55 | open IOObjectˢ public 56 | 57 | stack : ∀{A}{n : ℕ} (as : Vec A n) → Objectˢ (StackInterfaceˢ A) n 58 | objectMethod (stack as) (push a) = unit , stack (a ∷ as) 59 | objectMethod (stack (a ∷ as)) pop = a , stack as 60 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/VariableList.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.VariableList where 2 | 3 | open import Data.Product hiding (map) 4 | open import Data.List 5 | 6 | open import NativeIO 7 | open import StateSizedIO.GUI.WxBindingsFFI 8 | open import Relation.Binary.PropositionalEquality 9 | 10 | data VarList : Set₁ where 11 | [] : VarList 12 | addVar : (A : Set) → Var A → VarList → VarList 13 | 14 | 15 | prod : VarList → Set 16 | prod [] = Unit 17 | prod (addVar A v []) = A 18 | prod (addVar A v l) = A × prod l 19 | 20 | 21 | 22 | 23 | takeVar : (l : VarList) → NativeIO (prod l) 24 | takeVar [] = nativeReturn unit 25 | takeVar (addVar A v []) = nativeTakeVar {A} v 26 | takeVar (addVar A v (addVar B v′ l)) = 27 | nativeTakeVar {A} v native>>= λ a → 28 | takeVar (addVar B v′ l) native>>= λ rest → 29 | nativeReturn ( a , rest ) 30 | 31 | putVar : (l : VarList) → prod l → NativeIO Unit 32 | putVar [] _ = nativeReturn unit 33 | putVar (addVar A v []) a = nativePutVar {A} v a 34 | putVar (addVar A v (addVar B v′ l)) (a , rest) = 35 | nativePutVar {A} v a native>>= λ _ → 36 | putVar (addVar B v′ l) rest 37 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/VariableListForDispatchOnly.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.VariableListForDispatchOnly where 2 | 3 | open import Data.Product hiding (map) 4 | open import Data.List 5 | 6 | open import NativeIO 7 | open import StateSizedIO.GUI.WxBindingsFFI 8 | open import Relation.Binary.PropositionalEquality 9 | 10 | data VarList : Set₁ where 11 | [] : VarList 12 | addVar : (A : Set) → Var A → VarList → VarList 13 | 14 | prod : VarList → Set 15 | prod [] = Unit 16 | prod (addVar A v []) = A 17 | prod (addVar A v l) = A × prod l 18 | 19 | takeVar : (l : VarList) → NativeIO (prod l) 20 | takeVar [] = nativeReturn unit 21 | takeVar (addVar A v []) = 22 | nativeTakeVar {A} v native>>= λ a → nativeReturn a 23 | takeVar (addVar A v (addVar B v′ l)) = 24 | nativeTakeVar {A} v native>>= λ a → 25 | takeVar (addVar B v′ l) native>>= λ rest → nativeReturn ( a , rest ) 26 | 27 | putVar : (l : VarList) → prod l → NativeIO Unit 28 | putVar [] _ = nativeReturn unit 29 | putVar (addVar A v []) a = nativePutVar {A} v a 30 | putVar (addVar A v (addVar B v′ l)) (a , rest) = 31 | nativePutVar {A} v a native>>= λ _ → 32 | putVar (addVar B v′ l) rest native>>= nativeReturn 33 | 34 | 35 | dispatch : (l : VarList) → (prod l → NativeIO (prod l)) → NativeIO Unit 36 | dispatch l f = takeVar l native>>= λ a → 37 | f a native>>= λ a₁ → 38 | putVar l a₁ 39 | 40 | dispatchList : (l : VarList) → List (prod l → NativeIO (prod l)) → NativeIO Unit 41 | dispatchList l [] = nativeReturn unit 42 | dispatchList l (p ∷ rest) = dispatch l p native>>= λ _ → 43 | dispatchList l rest 44 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/WxGraphicsLib.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.WxGraphicsLib where 2 | 3 | open import SizedIO.Base 4 | open import StateSizedIO.GUI.BaseStateDependent 5 | 6 | open import Size 7 | 8 | open import Data.Bool.Base 9 | open import Data.List.Base 10 | 11 | open import Function 12 | 13 | open import NativeIO 14 | 15 | open import StateSizedIO.GUI.WxBindingsFFI 16 | open import StateSizedIO.GUI.VariableList 17 | 18 | 19 | data GuiLev1Command : Set where 20 | makeFrame : GuiLev1Command 21 | makeButton : Frame → GuiLev1Command 22 | addButton : Frame → Button → GuiLev1Command 23 | drawBitmap : DC → Bitmap → Point → Bool → GuiLev1Command 24 | repaint : Frame → GuiLev1Command 25 | 26 | GuiLev1Response : GuiLev1Command → Set 27 | GuiLev1Response makeFrame = Frame 28 | GuiLev1Response (makeButton _) = Button 29 | GuiLev1Response _ = Unit 30 | 31 | GuiLev1Interface : IOInterface 32 | Command GuiLev1Interface = GuiLev1Command 33 | Response GuiLev1Interface = GuiLev1Response 34 | 35 | 36 | GuiLev2State : Set₁ 37 | GuiLev2State = VarList 38 | 39 | 40 | data GuiLev2Command (s : GuiLev2State) : Set₁ where 41 | level1C : GuiLev1Command → GuiLev2Command s 42 | createVar : {A : Set} → A → GuiLev2Command s 43 | setButtonHandler : Button 44 | → List (prod s → IO GuiLev1Interface ∞ (prod s)) 45 | → GuiLev2Command s 46 | setOnPaint : Frame 47 | → List (prod s → DC → Rect → IO GuiLev1Interface ∞ (prod s)) 48 | → GuiLev2Command s 49 | 50 | 51 | GuiLev2Response : (s : GuiLev2State) → GuiLev2Command s → Set 52 | GuiLev2Response _ (level1C c) = GuiLev1Response c 53 | GuiLev2Response _ (createVar {A} a) = Var A 54 | GuiLev2Response _ _ = Unit 55 | 56 | 57 | GuiLev2Next : (s : GuiLev2State) → (c : GuiLev2Command s) 58 | → GuiLev2Response s c 59 | → GuiLev2State 60 | GuiLev2Next s (createVar {A} a) var = addVar A var s 61 | GuiLev2Next s _ _ = s 62 | 63 | 64 | GuiLev2Interface : IOInterfaceˢ 65 | Stateˢ GuiLev2Interface = GuiLev2State 66 | Commandˢ GuiLev2Interface = GuiLev2Command 67 | Responseˢ GuiLev2Interface = GuiLev2Response 68 | nextˢ GuiLev2Interface = GuiLev2Next 69 | 70 | 71 | translateLev1Local : (c : GuiLev1Command) → NativeIO (GuiLev1Response c) 72 | translateLev1Local makeFrame = nativeNewFrame "dummy title" 73 | translateLev1Local (makeButton fra) = nativeMakeButton fra "dummy button label" 74 | translateLev1Local (addButton fra bt) = nativeAddButton fra bt 75 | translateLev1Local (drawBitmap dc bm p b) = nativeDrawBitmap dc bm p b 76 | translateLev1Local (repaint fra) = nativeRepaint fra 77 | 78 | translateLev1 : {A : Set} → IO GuiLev1Interface ∞ A → NativeIO A 79 | translateLev1 = translateIO translateLev1Local 80 | 81 | translateLev1List : {A : Set} → List (IO GuiLev1Interface ∞ A) → List (NativeIO A) 82 | translateLev1List l = map translateLev1 l 83 | 84 | 85 | translateLev2Local : (s : GuiLev2State) 86 | → (c : GuiLev2Command s) 87 | → NativeIO (GuiLev2Response s c) 88 | translateLev2Local s (level1C c) = translateLev1Local c 89 | translateLev2Local s (createVar {A} a) = nativeNewVar {A} a 90 | translateLev2Local s (setButtonHandler bt proglist) = 91 | nativeSetButtonHandler bt 92 | (dispatchList s (map (λ prog → translateLev1 ∘ prog) proglist)) 93 | translateLev2Local s (setOnPaint fra proglist) = 94 | nativeSetOnPaint fra (λ dc rect → dispatchList s 95 | (map (λ prog aa → translateLev1 (prog aa dc rect)) proglist)) 96 | 97 | translateLev2 : ∀ {A s} → IOˢ GuiLev2Interface ∞ (λ _ → A) s → NativeIO A 98 | translateLev2 = translateIOˢ translateLev2Local 99 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/agdaCodeBrady.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module examplesPaperJFP.agdaCodeBrady where 3 | 4 | 5 | 6 | open import Data.List 7 | open import Agda.Builtin.Unit public renaming (⊤ to Unit; tt to triv) 8 | open import Data.Product 9 | open import examplesPaperJFP.StateDependentIO 10 | 11 | {- Brady's Effect -} 12 | 13 | Effect : Set₁ 14 | Effect = (Result : Set) → (InResource : Set) → (OutResource : Result → Set) → Set 15 | 16 | 17 | record MyEffect : Set₁ where 18 | field Ops : Set 19 | Result : Ops → Set 20 | InResource : Ops → Set 21 | OutResource : (o : Ops) → Result o → Set 22 | 23 | open MyEffect 24 | 25 | effectToIOInterfaceˢ : Effect → IOInterfaceˢ 26 | Stateˢ (effectToIOInterfaceˢ eff) = Set 27 | Commandˢ (effectToIOInterfaceˢ eff) s = 28 | Σ[ Result ∈ Set ] (Σ[ outR ∈ (Result → Set) ] (eff Result s outR)) 29 | Responseˢ (effectToIOInterfaceˢ eff) s (result , outR , op) = result 30 | nextˢ (effectToIOInterfaceˢ eff) s (result , outR , op) = outR 31 | 32 | const : { A B : Set} → B → A → B 33 | const b = λ _ → b 34 | 35 | data EFFECT′ : Set₁ where 36 | MkEff : Set → Effect → EFFECT′ 37 | 38 | EFFECT : Set₁ 39 | EFFECT = Set × Effect 40 | 41 | 42 | data State : Effect where 43 | Get : (a : Set) → State a a (λ _ → a) 44 | Put : (a : Set) → (b : Set) → State a a (λ _ → b) 45 | 46 | data myStateOps : Set₁ where 47 | get : Set → myStateOps 48 | put : Set → Set → myStateOps 49 | 50 | myState : MyEffect 51 | Ops myState = myStateOps 52 | Result myState (get a) = a 53 | InResource myState (get a) = a 54 | OutResource myState (get a) _ = a 55 | Result myState (put a b) = a 56 | InResource myState (put a b) = a 57 | OutResource myState (put a b) _ = b 58 | 59 | STATE : Set → EFFECT 60 | STATE x = ( x , State ) 61 | 62 | postulate String : Set 63 | 64 | data Stdio : Effect where 65 | PutStr : String → Stdio Unit Unit (const Unit) 66 | GetStr : Stdio String String (const String) 67 | 68 | STDIO : EFFECT 69 | STDIO = ( Unit , Stdio ) 70 | 71 | data Eff : (x : Set) → List EFFECT → (x → List EFFECT) → Set where 72 | get : (x : Set) → Eff x [ STATE x ] (const [ STATE x ]) 73 | put : (x : Set) → x → Eff Unit [ STATE x ] (const [ STATE x ]) 74 | putM : (x : Set) → (y : Set) → y → Eff Unit [ STATE x ] (const [ STATE y ]) 75 | update : (x : Set) → (x → x) → Eff Unit [ STATE x ] (const [ STATE x ]) 76 | 77 | 78 | data EffM : (m : Set → Set) → (res : Set) → 79 | (inEffects : List EFFECT) → 80 | (outEffects : res → List EFFECT) 81 | → Set where 82 | _>>=_ : {m : Set → Set} → {res : Set} → 83 | {inEffects : List EFFECT} → 84 | {outEffects : res → List EFFECT} → 85 | {res′ : Set} → 86 | {inEffects′ : res → List EFFECT} → 87 | {outEffects′ : res′ → List EFFECT} → 88 | EffM m res inEffects outEffects → 89 | ((x : res) → EffM m res′ (inEffects′ x) outEffects′) → 90 | EffM m res′ inEffects outEffects′ 91 | 92 | record SetInterfaceˢ : Set where 93 | field Commandˢ′ : Set → Set 94 | Responseˢ′ : (s : Set) → Commandˢ′ s → Set 95 | nextˢ′ : (s : Set) → (c : Commandˢ′ s) → Responseˢ′ s c → Set 96 | 97 | open SetInterfaceˢ 98 | 99 | module _ (I : SetInterfaceˢ ) (let Stateˢ = Set) 100 | (let C = Commandˢ′ I) (let R = Responseˢ′ I) 101 | (let next = nextˢ′ I) where 102 | handle : (M : Set → Set) → Set 103 | handle M = 104 | (A : Set) → (s : Stateˢ) → (c : C s) → (f : (r : R s c) → next s c r → M A) → M A 105 | 106 | postulate Char : Set 107 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/exampleFinFun.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.exampleFinFun where 2 | 3 | open import examplesPaperJFP.finn 4 | 5 | open import Data.Nat 6 | 7 | toℕ : ∀ {n} → Fin n → ℕ 8 | toℕ zero = 0 9 | toℕ (suc n) = suc (toℕ n) 10 | 11 | mutual 12 | _+e_ : ∀ {n m} → Even n → Even m → Even (n + m) 13 | 0p +e p = p 14 | sucp p +e q = sucp (p +o q) 15 | 16 | _+o_ : ∀ {n m} → Odd n → Even m → Odd (n + m) 17 | sucp p +o q = sucp (p +e q) 18 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/finn.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.finn where 2 | 3 | open import Data.Nat 4 | 5 | data Fin : ℕ → Set where 6 | zero : {n : ℕ} → Fin (suc n) 7 | suc : {n : ℕ} (i : Fin n) → Fin (suc n) 8 | 9 | mutual 10 | data Even : ℕ → Set where 11 | 0p : Even 0 12 | sucp : {n : ℕ} → Odd n → Even (suc n) 13 | 14 | data Odd : ℕ → Set where 15 | sucp : {n : ℕ} → Even n → Odd (suc n) 16 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/loadAllOOAgdaFilesAsInLibrary.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.loadAllOOAgdaFilesAsInLibrary where 2 | 3 | 4 | -- this file loads those files from the ooAgda paper 5 | -- which occur in the library as well 6 | -- the code as it occurred in the original paper after adapting to current Agda 7 | -- including any code not found directly inthe library 8 | -- can be found in loadAllOOAgdaPart1.agda 9 | -- and loadAllOOAgdaPart2.agda 10 | 11 | -- Source of paper: 12 | -- Andreas Abel, Stephan Adelsberger, Anton Setzer: 13 | -- Interactive Programming in Agda - Objects and Graphical User Interfaces 14 | -- Journal of Functional Programming, 27, 2017 15 | -- doi: http://dx.doi.org/10.1145/2976022.2976032 16 | -- authors copy: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.pdf 17 | -- bibtex: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.bib 18 | 19 | 20 | -- 1. Introdution 21 | -- 2. Introduction to Agda 22 | -- 3. Coalgebras in Agda 23 | -- 3.1. Coalgebra by example: Colists 24 | -- 3.2. Coalgebras in General 25 | -- 4. Interactive Programs in Agda 26 | -- 4.1. Interaction interfaces 27 | -- 4.2. Interaction trees 28 | -- 4.3. Running interactive programs 29 | -- 5. Objects in Agda 30 | 31 | -- code as in paper adapted to new Agda 32 | -- open import examplesPaperJFP.Object 33 | -- code as in library 34 | open import SizedIO.Object 35 | 36 | -- 6. Sized Coinductive Types 37 | -- 7. Interface Extension and Delegation 38 | 39 | -- code as in paper adapted to new Agda 40 | -- open import examplesPaperJFP.CounterCell 41 | -- code as in library 42 | open import Sized.CounterCell 43 | 44 | -- 8. State-Dependent Objects and IO 45 | -- 8.1 State-Dependent Interfaces 46 | -- 8.2 State-Dependent Objects 47 | -- 8.2.1. Example of Use of Safe Stack 48 | -- 8.3 Reasoning About Stateful Objects 49 | -- 8.3.1. Bisimilarity 50 | -- 8.3.2. Verifying stack laws} 51 | -- 8.3.3. Bisimilarity of different stack implementations 52 | -- 8.4. State-Dependent IO 53 | -- 9. A Drawing Program in Agda 54 | 55 | open import SizedIO.IOGraphicsLib 56 | 57 | 58 | -- 10. A Graphical User Interface using an Object 59 | -- 10.1. wxHaskell 60 | -- 10.2. A Library for Object-Based GUIs in Agda 61 | 62 | open import StateSizedIO.GUI.VariableList 63 | open import StateSizedIO.GUI.WxGraphicsLib 64 | 65 | -- 10.3 Example: A GUI controlling a Space Ship in Agda 66 | 67 | open import StateSized.GUI.SpaceShipSimpleVar 68 | open import StateSized.GUI.SpaceShipCell 69 | open import StateSized.GUI.SpaceShipAdvanced 70 | 71 | -- 11. Related Work 72 | -- 12. Conclusion 73 | -- Bibliography 74 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/loadAllOOAgdaPart1.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.loadAllOOAgdaPart1 where 2 | 3 | 4 | -- This file loads the files containing the code examples in the ooAgda paper 5 | -- ordered by sections 6 | -- The code is divided into files 7 | -- loadAllOOAgdaPart1.agda and loadAllOOAgdaPart2.agda 8 | -- loadAllOOAgdaPart1.agda loads code from Sect. 1 - 7 9 | -- loadAllOOAgdaPart2.agda loads code from Sect. 8 - 12 10 | -- 11 | -- The reason for splitting it into two files is that because of some 12 | -- Builtin the code from the first file cannot be loaded at the same 13 | -- time as the code from the second one. 14 | -- The difference is loading NativeIOSafe.agda vs NativeIO.agda 15 | -- 16 | -- Important changes due to updating it to current agda 17 | -- #################################################### 18 | -- * do is now a keyword and therefore replaced by exec 19 | -- * some other small changes mainly because of changes in the 20 | -- standard library have been made 21 | -- 22 | -- Files from the Library 23 | -- ######################### 24 | -- loadAllOOAgdaFilesAsInLibrary.agda 25 | -- loads corresponding code from the library which is slightly different 26 | -- from how it is represented in the paper 27 | 28 | 29 | -- Source of paper, presentations, slides 30 | -- ####################################### 31 | -- Andreas Abel, Stephan Adelsberger, Anton Setzer: 32 | -- Interactive Programming in Agda - Objects and Graphical User Interfaces 33 | -- Journal of Functional Programming, 27, 2017 34 | -- doi: http://dx.doi.org/10.1145/2976022.2976032 35 | -- authors copy: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.pdf 36 | -- bibtex: http://www.cs.swan.ac.uk/~csetzer/articles/ooAgda.bib 37 | 38 | 39 | -- 1. Introdution 40 | -- 2. Introduction to Agda 41 | 42 | import examplesPaperJFP.finn 43 | import examplesPaperJFP.exampleFinFun 44 | import examplesPaperJFP.Equality 45 | 46 | -- 3. Coalgebras in Agda 47 | -- 3.1. Coalgebra by example: Colists 48 | 49 | import examplesPaperJFP.Colists 50 | import examplesPaperJFP.Collatz 51 | 52 | -- 3.2. Coalgebras in General 53 | 54 | import examplesPaperJFP.Coalgebra 55 | 56 | -- 4. Interactive Programs in Agda 57 | -- 4.1. Interaction interfaces 58 | 59 | import examplesPaperJFP.NativeIOSafe 60 | import examplesPaperJFP.BasicIO 61 | import examplesPaperJFP.triangleRightOperator 62 | import examplesPaperJFP.ConsoleInterface 63 | 64 | -- 4.2. Interaction trees 65 | 66 | import examplesPaperJFP.Console 67 | import examplesPaperJFP.CatNonTerm 68 | import examplesPaperJFP.CatTerm 69 | 70 | -- 4.3. Running interactive programs 71 | 72 | import examplesPaperJFP.IOExampleConsole 73 | 74 | -- 5. Objects in Agda 75 | 76 | -- code as in paper adapted to new Agda 77 | import examplesPaperJFP.Object 78 | -- code as in library see loadAllOOAgdaFilesAsInLibrary.agda 79 | -- open import SizedIO.Object 80 | 81 | -- 6. Sized Coinductive Types 82 | 83 | import examplesPaperJFP.Sized 84 | 85 | -- 7. Interface Extension and Delegation 86 | 87 | -- code as in paper adapted to new Agda 88 | import examplesPaperJFP.CounterCell 89 | -- code as in library see loadAllOOAgdaFilesAsInLibrary.agda 90 | -- open import Sized.CounterCell 91 | 92 | -- *** Code from Sect 8 - 12 can be found in loadAllOOAgdaPart2.agda *** 93 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/loadAllOOAgdaPart2.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.loadAllOOAgdaPart2 where 2 | 3 | 4 | -- This is a continuation of the file loadAllOOAgdaPart1 5 | -- giving the code from the ooAgda paper 6 | -- This file was split into two because of a builtin IO which 7 | -- makes loading files from part1 and part2 incompatible. 8 | 9 | -- Note that some files which are directly in the libary can be found 10 | -- in loadAllOOAgdaFilesAsInLibrary 11 | 12 | -- Sect 1 - 7 are in loadAllOOAgdaPart1.agda 13 | 14 | -- 8. State-Dependent Objects and IO 15 | -- 8.1 State-Dependent Interfaces 16 | 17 | import examplesPaperJFP.StatefulObject 18 | 19 | -- 8.2 State-Dependent Objects 20 | -- 8.2.1. Example of Use of Safe Stack 21 | 22 | import examplesPaperJFP.safeFibStackMachineObjectOriented 23 | 24 | -- 8.3 Reasoning About Stateful Objects 25 | 26 | import examplesPaperJFP.StackBisim 27 | 28 | -- 8.3.1. Bisimilarity 29 | -- 8.3.2. Verifying stack laws} 30 | -- 8.3.3. Bisimilarity of different stack implementations 31 | 32 | 33 | -- 8.4. State-Dependent IO 34 | 35 | import examplesPaperJFP.StateDependentIO 36 | 37 | -- 9. A Drawing Program in Agda 38 | 39 | 40 | -- code as in paper adapted to new Agda 41 | open import examplesPaperJFP.IOGraphicsLib 42 | -- code as in library see loadAllOOAgdaFilesAsInLibrary.agda 43 | -- open import SizedIO.IOGraphicsLib 44 | open import examplesPaperJFP.ExampleDrawingProgram 45 | 46 | -- 10. A Graphical User Interface using an Object 47 | -- 10.1. wxHaskell 48 | -- 10.2. A Library for Object-Based GUIs in Agda 49 | 50 | open import examplesPaperJFP.VariableList 51 | open import examplesPaperJFP.WxGraphicsLib 52 | open import examplesPaperJFP.VariableListForDispatchOnly 53 | 54 | -- 10.3 Example: A GUI controlling a Space Ship in Agda 55 | 56 | open import examplesPaperJFP.SpaceShipSimpleVar 57 | open import examplesPaperJFP.SpaceShipCell 58 | open import examplesPaperJFP.SpaceShipAdvanced 59 | 60 | -- 11. Related Work 61 | 62 | open import examplesPaperJFP.agdaCodeBrady 63 | 64 | -- 12. Conclusion 65 | -- Bibliography 66 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/safeFibStackMachineObjectOriented.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.safeFibStackMachineObjectOriented where 2 | 3 | open import Data.Nat 4 | open import Data.List 5 | open import Data.Vec 6 | open import Data.Sum 7 | open import Data.Fin renaming (_+_ to _+f_) 8 | open import Data.Product 9 | open import examplesPaperJFP.StatefulObject 10 | open import examplesPaperJFP.StackBisim 11 | 12 | open import examplesPaperJFP.triangleRightOperator 13 | 14 | 15 | 16 | {- Object based version of safe stack 17 | it is essentially the last part of safeFibStackMachine.agda 18 | with the names simplified -} 19 | 20 | 21 | {- the state is what we are computing right now: 22 | fib n menas we need to compute fib n 23 | val k means we have computed the value k 24 | -} 25 | 26 | 27 | 28 | data FibState : Set where 29 | fib : ℕ → FibState 30 | val : ℕ → FibState 31 | 32 | data FibStackEl : Set where 33 | _+• : ℕ → FibStackEl 34 | •+fib_ : ℕ → FibStackEl 35 | 36 | FibStack : ℕ → Set 37 | FibStack = Objectˢ (StackInterfaceˢ FibStackEl) 38 | 39 | FibStackmachine : Set 40 | FibStackmachine = Σ[ n ∈ ℕ ] (FibState × FibStack n) 41 | 42 | reduce : FibStackmachine → FibStackmachine ⊎ ℕ 43 | reduce (n , fib 0 , stack) = inj₁ (n , val 1 , stack) 44 | reduce (n , fib 1 , stack) = inj₁ (n , val 1 , stack) 45 | reduce (n , fib (suc (suc m)) , stack) = 46 | objectMethod stack (push (•+fib m)) ▹ λ { (_ , stack₁) → 47 | inj₁ ( suc n , fib (suc m) , stack₁) } 48 | reduce (0 , val k , stack) = inj₂ k 49 | reduce (suc n , val k , stack) = 50 | objectMethod stack pop ▹ λ { (k′ +• , stack₁) → 51 | inj₁ (n , val (k′ + k) , stack₁) 52 | ; (•+fib m , stack₁) → 53 | objectMethod stack₁ (push (k +•)) ▹ λ { (_ , stack₂) → 54 | inj₁ (suc n , fib m , stack₂) }} 55 | 56 | iter : (m : ℕ) → FibStackmachine → FibStackmachine ⊎ ℕ 57 | iter 0 s = inj₁ s 58 | iter (suc n) s with reduce s 59 | ... | inj₁ s′ = iter n s′ 60 | ... | inj₂ m = inj₂ m 61 | 62 | computeFib : ℕ → ℕ → FibStackmachine ⊎ ℕ 63 | computeFib n m = iter n (0 , fib m , stack []) 64 | 65 | fibO0 : FibStackmachine ⊎ ℕ 66 | fibO0 = computeFib 2 0 67 | 68 | -- evaluates to inj₂ 1 69 | 70 | fibO1 : FibStackmachine ⊎ ℕ 71 | fibO1 = computeFib 2 1 72 | -- evaluates to inj₂ 1 73 | 74 | fibO2 : FibStackmachine ⊎ ℕ 75 | fibO2 = computeFib 10 2 76 | -- evaluates to inj₂ 2 77 | 78 | fibO3 : FibStackmachine ⊎ ℕ 79 | fibO3 = computeFib 14 3 80 | -- evaluates to inj₂ 3 81 | 82 | fibO4 : FibStackmachine ⊎ ℕ 83 | fibO4 = computeFib 30 4 84 | -- evaluates to inj₂ 5 85 | 86 | fibO5 : FibStackmachine ⊎ ℕ 87 | fibO5 = computeFib 30 5 88 | -- evaluates to inj₂ 8 89 | 90 | {-# TERMINATING #-} 91 | computeFibRec : FibStackmachine → ℕ 92 | computeFibRec s with reduce s 93 | ... | inj₁ s′ = computeFibRec s′ 94 | ... | inj₂ k = k 95 | 96 | 97 | fibUsingStack : ℕ → ℕ 98 | fibUsingStack m = computeFibRec (0 , fib m , stack []) 99 | 100 | 101 | test : List ℕ 102 | test = fibUsingStack 0 ∷ fibUsingStack 1 ∷ fibUsingStack 2 ∷ fibUsingStack 3 ∷ fibUsingStack 4 ∷ fibUsingStack 5 ∷ fibUsingStack 6 ∷ [] 103 | 104 | testExpected : List ℕ 105 | testExpected = 1 ∷ 1 ∷ 2 ∷ 3 ∷ 5 ∷ 8 ∷ 13 ∷ [] 106 | -------------------------------------------------------------------------------- /examples/examplesPaperJFP/triangleRightOperator.agda: -------------------------------------------------------------------------------- 1 | module examplesPaperJFP.triangleRightOperator where 2 | 3 | _▹_ : ∀{A B : Set} → A → (A → B) → B 4 | a ▹ f = f a 5 | -------------------------------------------------------------------------------- /examples/java/Cell.java: -------------------------------------------------------------------------------- 1 | interface Cell { 2 | void put (A s); 3 | A get (); 4 | } 5 | -------------------------------------------------------------------------------- /examples/java/CounterCell.java: -------------------------------------------------------------------------------- 1 | class CounterCell implements StatsCell { 2 | 3 | Cell cell; 4 | int ngets, nputs; 5 | 6 | CounterCell (Cell c, int g, int p) { 7 | cell = c; 8 | ngets = g; 9 | nputs = p; 10 | } 11 | 12 | public A get() { 13 | ngets++; 14 | return cell.get(); 15 | } 16 | 17 | public void put (A s) { 18 | nputs++; 19 | cell.put(s); 20 | } 21 | 22 | public void stats() { 23 | System.out.println ("Counted " 24 | + ngets + " calls to get and " 25 | + nputs + " calls to put."); 26 | } 27 | 28 | public static void program (String arg) { 29 | CounterCell c = new CounterCell(new SimpleCell("Start"), 0, 0); 30 | String s = c.get(); 31 | System.out.println(s); 32 | c.put(arg); 33 | s = c.get(); 34 | System.out.println(s); 35 | c.put("Over!"); 36 | c.stats(); 37 | return; 38 | } 39 | 40 | public static void main (String[] args) { 41 | program ("Hello"); 42 | } 43 | } 44 | -------------------------------------------------------------------------------- /examples/java/Makefile: -------------------------------------------------------------------------------- 1 | default : Cell.class SimpleCell CounterCell Parrot SelfRefExample 2 | 3 | % : %.class 4 | @echo "======================================================================" 5 | @echo "$*" 6 | @echo "======================================================================" 7 | @java $* Hello world 8 | 9 | %.class : %.java 10 | javac $< 11 | 12 | # EOF 13 | -------------------------------------------------------------------------------- /examples/java/Parrot.java: -------------------------------------------------------------------------------- 1 | class Parrot implements Cell { 2 | 3 | Cell cell; 4 | 5 | Parrot (Cell c) { cell = c; } 6 | 7 | public String get() { 8 | return "(" + cell.get() + ") is what parrot got"; 9 | } 10 | public void put (String s) { 11 | cell.put("parrot puts (" + s + ")"); 12 | } 13 | 14 | public static void main (String[] args) { 15 | Parrot c = new Parrot(new SimpleCell("Start")); 16 | String s = c.get(); 17 | System.out.println(s); 18 | c.put(s); 19 | s = c.get(); 20 | System.out.println(s); 21 | } 22 | } 23 | -------------------------------------------------------------------------------- /examples/java/SelfRefExample.java: -------------------------------------------------------------------------------- 1 | public class SelfRefExample { 2 | 3 | public void m1() { 4 | System.out.println("m1 called"); 5 | m2(); // self ref call 6 | } 7 | 8 | public void m2() { 9 | System.out.println("m2 called"); 10 | } 11 | 12 | public static void main(String[] args) { 13 | SelfRefExample ex = new SelfRefExample(); 14 | ex.m1(); 15 | } 16 | } 17 | -------------------------------------------------------------------------------- /examples/java/SimpleCell.java: -------------------------------------------------------------------------------- 1 | class SimpleCell implements Cell { 2 | 3 | A content; 4 | 5 | SimpleCell (A s) { content = s; } 6 | 7 | public void put (A s) { 8 | System.err.println("putting(" + s + ")"); 9 | content = s; 10 | } 11 | 12 | public A get () { 13 | System.err.println("getting(" + content + ")"); 14 | return content; 15 | } 16 | 17 | public static void program () { 18 | SimpleCell c = new SimpleCell("Start"); 19 | String s = System.console().readLine(); 20 | if (s == null) return; else { 21 | c.put(s); 22 | s = c.get(); 23 | System.out.println(s); 24 | program(); 25 | } 26 | } 27 | 28 | public static void main (String[] args) { 29 | program(); 30 | } 31 | } 32 | -------------------------------------------------------------------------------- /examples/java/Stack/README: -------------------------------------------------------------------------------- 1 | 2 | 3 | modelt after java.util Interface Deque 4 | 5 | pop explicitly throws a 'java.util.NoSuchElementException' exception, while the java docs of java.util Interface Deque don't use a checked exception; the java docs use a runtime exception. 6 | 7 | -------------------------------------------------------------------------------- /examples/java/Stack/Stack.java: -------------------------------------------------------------------------------- 1 | public interface Stack { 2 | void push(E e); 3 | 4 | /** @throws EmptyStackException if the stack is empty **/ 5 | E pop() throws java.util.EmptyStackException; 6 | } 7 | -------------------------------------------------------------------------------- /examples/java/Stack/StackLinkedList.java: -------------------------------------------------------------------------------- 1 | import java.util.LinkedList; 2 | 3 | public class StackLinkedList implements Stack { 4 | private LinkedList list = new LinkedList(); 5 | 6 | public void push(T item) { list.addFirst(item); } 7 | 8 | public T pop() throws java.util.NoSuchElementException 9 | { 10 | if (list.isEmpty()) { throw new java.util.NoSuchElementException(); } 11 | else return list.removeFirst(); 12 | } 13 | 14 | 15 | public static void main(String[] args) { 16 | StackLinkedList s = new StackLinkedList(); 17 | 18 | String str1 = System.console().readLine(); 19 | s.push(str1); 20 | String str2 = System.console().readLine(); 21 | s.push(str2); 22 | 23 | System.out.println("first pop: " + s.pop()); 24 | 25 | String str3 = System.console().readLine(); 26 | s.push(str3); 27 | 28 | System.out.println("second pop: " + s.pop()); 29 | 30 | System.out.println("third pop: " + s.pop()); 31 | System.out.println("fourth pop: " + s.pop()); 32 | } 33 | } 34 | -------------------------------------------------------------------------------- /examples/java/StatsCell.java: -------------------------------------------------------------------------------- 1 | interface StatsCell extends Cell { 2 | void stats(); 3 | } 4 | -------------------------------------------------------------------------------- /examples/ooAgdaExamples.agda-lib: -------------------------------------------------------------------------------- 1 | name: ooAgdaExamples 2 | depend: standard-library 3 | include: . ../src 4 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/NativeIOSafe.agda: -------------------------------------------------------------------------------- 1 | module NativeIOSafe where 2 | 3 | open import Data.Maybe.Base using (Maybe; nothing; just) public 4 | open import Data.String.Base using (String) public 5 | 6 | record Unit : Set where 7 | constructor unit 8 | 9 | {-# COMPILED_DATA Unit () () #-} 10 | 11 | postulate 12 | NativeIO : Set → Set 13 | nativeReturn : {A : Set} → A → NativeIO A 14 | _native>>=_ : {A B : Set} → NativeIO A → (A → NativeIO B) → NativeIO B 15 | 16 | {-# BUILTIN IO NativeIO #-} 17 | {-# COMPILED_TYPE NativeIO IO #-} -- IO.FFI.AgdaIO 18 | {-# COMPILED _native>>=_ (\_ _ -> (>>=) :: IO a -> (a -> IO b) -> IO b) #-} 19 | {-# COMPILED nativeReturn (\_ -> return :: a -> IO a) #-} 20 | 21 | postulate 22 | nativeGetLine : NativeIO (Maybe String) 23 | nativePutStrLn : String → NativeIO Unit 24 | 25 | {-# IMPORT Data.Text #-} 26 | {-# IMPORT System.IO.Error #-} 27 | 28 | {-# COMPILED nativePutStrLn (\ s -> putStrLn (Data.Text.unpack s)) #-} 29 | {-# COMPILED nativeGetLine (fmap (Just . Data.Text.pack) getLine `System.IO.Error.catchIOError` \ err -> if System.IO.Error.isEOFError err then return Nothing else ioError err) #-} 30 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/StateSized/GUI/ship.ico: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/ooAgda/7cc45e0148a4a508d20ed67e791544c30fecd795/presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/StateSized/GUI/ship.ico -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/drawingProgram.agda: -------------------------------------------------------------------------------- 1 | module drawingProgram where 2 | 3 | open import Unit 4 | open import Data.Bool.Base 5 | open import Data.Char.Base renaming (primCharEquality to charEquality) 6 | open import Data.Nat.Base hiding (_≟_;_⊓_; _+_; _*_) 7 | open import Data.List.Base hiding (_++_) 8 | open import Data.Integer.Base hiding (suc) 9 | open import Data.String.Base 10 | open import Data.Maybe.Base 11 | open import NativeIO 12 | 13 | open import Function 14 | 15 | open import Size renaming (Size to AgdaSize) 16 | 17 | open import SizedIO.Base 18 | open import SizedIO.IOGraphicsLib hiding (GraphicsCommands; GraphicsResponses; GraphicsInterface; IOGraphics; getWindowEvent; openWindow; drawInWindow; closeWindow; translateIOGraphicsLocal; translateIOGraphics) 19 | 20 | open import NativeInt 21 | -- open import NativeIOSafe 22 | 23 | 24 | data GraphicsCommands : Set where 25 | closeWindow : Window → GraphicsCommands 26 | maybeGetWindowEvent : Window → GraphicsCommands 27 | getWindowEvent : Window → GraphicsCommands 28 | openWindowNotEx : String → Size → GraphicsCommands 29 | openWindow : String → (Maybe Point) → (Maybe Size) 30 | → RedrawMode → (Maybe Word32) → GraphicsCommands 31 | timeGetTime : GraphicsCommands 32 | drawInWindow : Window → Graphic → GraphicsCommands 33 | print : String → GraphicsCommands 34 | 35 | 36 | GraphicsResponses : GraphicsCommands → Set 37 | GraphicsResponses (maybeGetWindowEvent w) = Maybe Event 38 | GraphicsResponses (getWindowEvent w) = Event 39 | GraphicsResponses (closeWindow w) = Unit 40 | GraphicsResponses (openWindowNotEx s s') = Window 41 | GraphicsResponses (openWindow s p s' r w) = Window 42 | GraphicsResponses timeGetTime = Word32 43 | GraphicsResponses _ = Unit 44 | 45 | 46 | GraphicsInterface : IOInterface 47 | Command GraphicsInterface = GraphicsCommands 48 | Response GraphicsInterface = GraphicsResponses 49 | 50 | IOGraphics : AgdaSize → Set → Set 51 | IOGraphics i = IO GraphicsInterface i 52 | 53 | 54 | translateIOGraphicsLocal : (c : GraphicsCommands) → NativeIO (GraphicsResponses c) 55 | translateIOGraphicsLocal (maybeGetWindowEvent w) = nativeMaybeGetWindowEvent w 56 | translateIOGraphicsLocal (getWindowEvent w) = nativeGetWindowEvent w 57 | translateIOGraphicsLocal (closeWindow w) = nativeCloseWindow w 58 | translateIOGraphicsLocal (openWindowNotEx str size) = nativeOpenWindow str size 59 | translateIOGraphicsLocal (openWindow str point size mode word) = nativeOpenWindowEx str point size mode word 60 | translateIOGraphicsLocal timeGetTime = nativeTimeGetTime 61 | translateIOGraphicsLocal (drawInWindow w g) = nativeDrawInWindow w g 62 | translateIOGraphicsLocal (print s) = nativePutStrLn s 63 | 64 | 65 | translateIOGraphics : {A : Set} → IOGraphics ∞ A → NativeIO A 66 | translateIOGraphics = translateIO translateIOGraphicsLocal 67 | 68 | 69 | integerLess : ℤ → ℤ → Bool 70 | integerLess x y with ∣(y - (x ⊓ y))∣ 71 | ... | zero = true 72 | ... | z = false 73 | 74 | line : Point → Point → Graphic 75 | line p newpoint = withColor red (polygon 76 | (nativePoint x y 77 | ∷ nativePoint a b 78 | ∷ nativePoint (a + xAddition) (b + yAddition) 79 | ∷ nativePoint (x + xAddition) (y + yAddition) 80 | ∷ [] ) ) 81 | where 82 | x = nativeProj1Point p 83 | y = nativeProj2Point p 84 | a = nativeProj1Point newpoint 85 | b = nativeProj2Point newpoint 86 | 87 | diffx = + ∣ (a - x) ∣ 88 | diffy = + ∣ (b - y) ∣ 89 | 90 | diffx*3 = diffx * (+ 3) 91 | diffy*3 = diffy * (+ 3) 92 | 93 | condition = (integerLess diffx diffy*3) ∧ (integerLess diffy diffx*3) 94 | 95 | xAddition = if condition then + 2 else + 1 96 | yAddition = if condition then + 2 else + 1 97 | 98 | 99 | 100 | State = Maybe Point 101 | 102 | loop : ∀{i} → Window → State → IOGraphics i Unit 103 | force (loop w s) = do' (getWindowEvent w) 104 | λ{ (Key c t) → if charEquality c 'x' then do (closeWindow w) return 105 | else loop w s 106 | ; (MouseMove p₂) → case s of 107 | λ{ nothing → loop w (just p₂) ; 108 | (just p₁) → do (drawInWindow w (line p₁ p₂)) λ _ → 109 | loop w (just p₂) } 110 | ; _ → loop w s } 111 | 112 | 113 | program : ∀{i} → IOGraphics i Unit 114 | program = 115 | do (openWindow "Drawing Program" nothing (just (size (+ 1000) (+ 1000))) 116 | nativeDrawGraphic nothing) λ window → 117 | loop window nothing 118 | 119 | main : NativeIO Unit 120 | main = nativeRunGraphics (translateIOGraphics program) 121 | 122 | 123 | 124 | 125 | 126 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/interactiveProgramsAgda.agda: -------------------------------------------------------------------------------- 1 | module interactiveProgramsAgda where 2 | 3 | open import NativeIO 4 | open import Function 5 | open import Size 6 | open import Data.Maybe.Base 7 | 8 | record IOInterface : Set₁ where 9 | field Command : Set 10 | Response : (c : Command) → Set 11 | 12 | open IOInterface public 13 | 14 | mutual 15 | record IO (I : IOInterface) (i : Size) (A : Set) : Set where 16 | coinductive 17 | constructor delay 18 | field force : {j : Size< i} → IO' I j A 19 | 20 | data IO' (I : IOInterface) (i : Size) (A : Set) : Set where 21 | do' : (c : Command I) (f : Response I c → IO I i A) → IO' I i A 22 | return' : (a : A) → IO' I i A 23 | 24 | open IO public 25 | 26 | 27 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) where 28 | do : ∀{i A} (c : C) (f : R c → IO I i A) → IO I i A 29 | force (do c f) = do' c f 30 | 31 | return : ∀{i A} (a : A) → IO I i A 32 | force (return a) = return' a 33 | 34 | infixl 2 _>>=_ 35 | 36 | _>>=_ : ∀{i A B} (m : IO I i A) (k : A → IO I i B) → IO I i B 37 | force (m >>= k) with force m 38 | ... | do' c f = do' c λ x → f x >>= k 39 | ... | return' a = force (k a) 40 | 41 | 42 | {-# NON_TERMINATING #-} 43 | translateIO : ∀ {A} (tr : (c : C) → NativeIO (R c)) → IO I ∞ A → NativeIO A 44 | translateIO tr m = case (force m) of λ 45 | { (do' c f) → (tr c) native>>= λ r → translateIO tr (f r) 46 | ; (return' a) → nativeReturn a 47 | } 48 | 49 | 50 | {- Specific case of Console IO -} 51 | 52 | data ConsoleCommand : Set where 53 | getLine : ConsoleCommand 54 | putStrLn : String → ConsoleCommand 55 | 56 | ConsoleResponse : ConsoleCommand → Set 57 | ConsoleResponse getLine = String 58 | ConsoleResponse (putStrLn s) = Unit 59 | 60 | ConsoleInterface : IOInterface 61 | Command ConsoleInterface = ConsoleCommand 62 | Response ConsoleInterface = ConsoleResponse 63 | 64 | IOConsole : (i : Size) → Set → Set 65 | IOConsole i = IO ConsoleInterface i 66 | 67 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 68 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 69 | translateIOConsoleLocal getLine = nativeGetLine 70 | 71 | translateIOConsole : {A : Set} → IOConsole ∞ A → NativeIO A 72 | translateIOConsole = translateIO translateIOConsoleLocal 73 | 74 | cat : ∀ {i} → IOConsole i Unit 75 | force cat = do' getLine λ line → 76 | do (putStrLn line) λ _ → 77 | cat 78 | 79 | 80 | main : NativeIO Unit 81 | main = translateIOConsole cat 82 | 83 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/interactiveProgramsAgda.old.agda: -------------------------------------------------------------------------------- 1 | module interactiveProgramsAgda where 2 | 3 | open import NativeIOSafe 4 | open import Function 5 | open import Size 6 | 7 | record IOInterface : Set₁ where 8 | field Command : Set 9 | Response : (c : Command) → Set 10 | 11 | open IOInterface public 12 | 13 | mutual 14 | record IO (I : IOInterface) (i : Size) (A : Set) : Set where 15 | coinductive 16 | constructor delay 17 | field force : {j : Size< i} → IO' I j A 18 | 19 | data IO' (I : IOInterface) (i : Size) (A : Set) : Set where 20 | do' : (c : Command I) (f : Response I c → IO I i A) → IO' I i A 21 | return' : (a : A) → IO' I i A 22 | 23 | open IO public 24 | 25 | {- 26 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) where 27 | do : ∀{A} (c : C) (f : R c → IO I A) → IO I A 28 | force (do c f) = do' c f 29 | 30 | return : ∀{A} (a : A) → IO I A 31 | force (return a) = return' a 32 | 33 | infixl 2 _>>=_ 34 | 35 | _>>=_ : ∀{A B} (m : IO I A) (k : A → IO I B) → IO I B 36 | force (m >>= k) with force m 37 | ... | do' c f = do' c λ x → f x >>= k 38 | ... | return' a = force (k a) 39 | 40 | 41 | {-# NON_TERMINATING #-} 42 | translateIO : ∀ {A} (tr : (c : C) → NativeIO (R c)) → IO I A → NativeIO A 43 | translateIO tr m = case (force m) of λ 44 | { (do' c f) → (tr c) native>>= λ r → translateIO tr (f r) 45 | ; (return' a) → nativeReturn a 46 | } 47 | 48 | 49 | data ConsoleCommand : Set where 50 | getLine : ConsoleCommand 51 | putStrLn : String → ConsoleCommand 52 | 53 | ConsoleResponse : ConsoleCommand → Set 54 | ConsoleResponse getLine = Maybe String 55 | ConsoleResponse (putStrLn s) = Unit 56 | 57 | ConsoleInterface : IOInterface 58 | Command ConsoleInterface = ConsoleCommand 59 | Response ConsoleInterface = ConsoleResponse 60 | 61 | 62 | IOConsole : Set → Set 63 | IOConsole = IO ConsoleInterface 64 | 65 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 66 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 67 | translateIOConsoleLocal getLine = nativeGetLine 68 | 69 | translateIOConsole : {A : Set} → IOConsole A → NativeIO A 70 | translateIOConsole = translateIO translateIOConsoleLocal 71 | 72 | 73 | {-# NON_TERMINATING #-} 74 | cat : IOConsole Unit 75 | force cat = do' getLine λ{ nothing → return _ ; 76 | (just line) → 77 | do (putStrLn line) λ _ → 78 | cat 79 | } 80 | 81 | mutual 82 | cat' : IOConsole Unit 83 | force cat' = do' getLine λ{ nothing → return _ ; 84 | (just line) → 85 | cat'' line} 86 | 87 | cat'' : String → IOConsole Unit 88 | force (cat'' line) = do' (putStrLn line) λ _ → 89 | cat' 90 | 91 | 92 | main : NativeIO Unit 93 | main = translateIOConsole cat 94 | -} 95 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/interactiveProgramsAgdaUnsized.agda: -------------------------------------------------------------------------------- 1 | module interactiveProgramsAgdaUnsized where 2 | 3 | open import NativeIO 4 | open import Function 5 | 6 | 7 | record IOInterface : Set₁ where 8 | field Command : Set 9 | Response : (c : Command) → Set 10 | 11 | open IOInterface public 12 | 13 | mutual 14 | record IO (I : IOInterface) (A : Set) : Set where 15 | coinductive 16 | constructor delay 17 | field force : IO' I A 18 | 19 | data IO' (I : IOInterface) (A : Set) : Set where 20 | do' : (c : Command I) (f : Response I c → IO I A) → IO' I A 21 | return' : (a : A) → IO' I A 22 | 23 | open IO public 24 | 25 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) where 26 | do : ∀{A} (c : C) (f : R c → IO I A) → IO I A 27 | force (do c f) = do' c f 28 | 29 | return : ∀{A} (a : A) → IO I A 30 | force (return a) = return' a 31 | 32 | infixl 2 _>>=_ 33 | 34 | _>>=_ : ∀{A B} (m : IO I A) (k : A → IO I B) → IO I B 35 | force (m >>= k) with force m 36 | ... | do' c f = do' c λ x → f x >>= k 37 | ... | return' a = force (k a) 38 | 39 | 40 | {-# NON_TERMINATING #-} 41 | translateIO : ∀ {A} (tr : (c : C) → NativeIO (R c)) → IO I A → NativeIO A 42 | translateIO tr m = case (force m) of λ 43 | { (do' c f) → (tr c) native>>= λ r → translateIO tr (f r) 44 | ; (return' a) → nativeReturn a 45 | } 46 | 47 | 48 | data ConsoleCommand : Set where 49 | getLine : ConsoleCommand 50 | putStrLn : String → ConsoleCommand 51 | 52 | ConsoleResponse : ConsoleCommand → Set 53 | ConsoleResponse getLine = String 54 | ConsoleResponse (putStrLn s) = Unit 55 | 56 | ConsoleInterface : IOInterface 57 | Command ConsoleInterface = ConsoleCommand 58 | Response ConsoleInterface = ConsoleResponse 59 | 60 | 61 | IOConsole : Set → Set 62 | IOConsole = IO ConsoleInterface 63 | 64 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 65 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 66 | translateIOConsoleLocal getLine = nativeGetLine 67 | 68 | translateIOConsole : {A : Set} → IOConsole A → NativeIO A 69 | translateIOConsole = translateIO translateIOConsoleLocal 70 | 71 | 72 | {-# NON_TERMINATING #-} 73 | cat : IOConsole Unit 74 | force cat = do' getLine λ{ line → 75 | do (putStrLn line) λ _ → 76 | cat 77 | } 78 | 79 | mutual 80 | cat' : IOConsole Unit 81 | force cat' = do' getLine λ{ line → 82 | cat'' line} 83 | 84 | cat'' : String → IOConsole Unit 85 | force (cat'' line) = do' (putStrLn line) λ _ → 86 | cat' 87 | 88 | 89 | main : NativeIO Unit 90 | main = translateIOConsole cat 91 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/interfaceExtensionAndDelegation.agda: -------------------------------------------------------------------------------- 1 | module interfaceExtensionAndDelegation where 2 | 3 | open import Data.Product 4 | open import Data.Nat.Base 5 | open import Data.Nat.Show 6 | open import Data.String.Base using (String; _++_) 7 | open import Size 8 | open import NativeIO 9 | open import interactiveProgramsAgda using (ConsoleInterface; _>>=_; do; 10 | IO; return; putStrLn; 11 | translateIOConsole ) 12 | open import objectsInAgda using (Interface; Method; Result; CellMethod; 13 | get; put; CellResult; cellI; IOObject; 14 | CellC; method; simpleCell ) 15 | 16 | 17 | 18 | data CounterMethod A : Set where 19 | super : (m : CellMethod A) → CounterMethod A 20 | stats : CounterMethod A 21 | 22 | statsCellI : (A : Set) → Interface 23 | Method (statsCellI A) = CounterMethod A 24 | Result (statsCellI A) (super m) = Result (cellI A) m 25 | Result (statsCellI A) stats = Unit 26 | 27 | CounterC : (i : Size) → Set 28 | CounterC = IOObject ConsoleInterface (statsCellI String) 29 | 30 | pattern getᶜ = super get 31 | pattern putᶜ x = super (put x) 32 | 33 | {- Methods of CounterC are now 34 | getᶜ (putᶜ x) stats 35 | -} 36 | 37 | 38 | counterCell : ∀{i} (c : CellC i) (ngets nputs : ℕ) → CounterC i 39 | method (counterCell c ngets nputs) getᶜ = 40 | method c get >>= λ { (s , c') → 41 | return (s , counterCell c' (1 + ngets) nputs) } 42 | 43 | method (counterCell c ngets nputs) (putᶜ x) = 44 | method c (put x) >>= λ { (_ , c') → 45 | return (_ , counterCell c' ngets (1 + nputs)) } 46 | 47 | method (counterCell c ngets nputs) stats = 48 | do (putStrLn ("Counted " 49 | ++ show ngets ++ " calls to get and " 50 | ++ show nputs ++ " calls to put.")) λ _ → 51 | return (_ , counterCell c ngets nputs) 52 | 53 | program : String → IO ConsoleInterface ∞ Unit 54 | program arg = 55 | let c₀ = counterCell (simpleCell "Start") 0 0 in 56 | method c₀ getᶜ >>= λ{ (s , c₁) → 57 | do (putStrLn s) λ _ → 58 | method c₁ (putᶜ arg) >>= λ{ (_ , c₂) → 59 | method c₂ getᶜ >>= λ{ (s' , c₃) → 60 | do (putStrLn s') λ _ → 61 | method c₃ (putᶜ "Over!") >>= λ{ (_ , c₄) → 62 | method c₄ stats >>= λ{ (_ , c₅) → 63 | return _ }}}}} 64 | 65 | 66 | main : NativeIO Unit 67 | main = translateIOConsole (program "Hello") 68 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/main.agda: -------------------------------------------------------------------------------- 1 | module main where 2 | 3 | 4 | {- unsized Versions omitted in talk -} 5 | import interactiveProgramsAgdaUnsized 6 | import objectsInAgdaUnsized 7 | {- end unsized Versions omitted in talk -} 8 | 9 | 10 | import interactiveProgramsAgda 11 | import objectsInAgda 12 | 13 | {- Probably omit because of time -} 14 | import interfaceExtensionAndDelegation 15 | {- End Probably omit because of time -} 16 | 17 | 18 | {- Omit later parts about proofs -} 19 | import stateDependentObjects 20 | {- End Omit later parts about proofs -} 21 | 22 | import stateDependentIO 23 | 24 | import drawingProgram 25 | 26 | import objectOrientedGui 27 | import spaceShipSimpleVar 28 | import spaceShipCell 29 | import spaceShipAdvanced 30 | 31 | 32 | 33 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/objectsInAgda.agda: -------------------------------------------------------------------------------- 1 | module objectsInAgda where 2 | 3 | open import Data.Product 4 | open import interactiveProgramsAgda hiding (main) 5 | open import NativeIO 6 | open import Data.String.Base 7 | 8 | 9 | record Interface : Set₁ where 10 | field Method : Set 11 | Result : (m : Method) → Set 12 | 13 | open Interface public 14 | 15 | record Object (I : Interface) : Set where 16 | coinductive 17 | field objectMethod : (m : Method I) → Result I m × Object I 18 | 19 | open Object public 20 | open import Size 21 | 22 | 23 | record IOObject (Iᵢₒ : IOInterface) (I : Interface) (i : Size) : Set where 24 | coinductive 25 | field method : ∀{j : Size< i} (m : Method I) 26 | → IO Iᵢₒ ∞ (Result I m × IOObject Iᵢₒ I j) 27 | 28 | open IOObject public 29 | 30 | {- Example Cell -} 31 | 32 | data CellMethod A : Set where 33 | get : CellMethod A 34 | put : A → CellMethod A 35 | 36 | CellResult : ∀{A} → CellMethod A → Set 37 | CellResult {A} get = A 38 | CellResult (put _) = Unit 39 | 40 | cellI : (A : Set) → Interface 41 | Method (cellI A) = CellMethod A 42 | Result (cellI A) m = CellResult m 43 | 44 | 45 | {- A cell as an Object -} 46 | 47 | Cell : Set → Set 48 | Cell A = Object (cellI A) 49 | 50 | cell : {A : Set} → A → Cell A 51 | objectMethod (cell a) get = a , cell a 52 | objectMethod (cell a) (put a') = unit , cell a' 53 | 54 | 55 | 56 | 57 | {- A cell as an IOObject -} 58 | 59 | CellC : (i : Size) → Set 60 | CellC = IOObject ConsoleInterface (cellI String) 61 | 62 | 63 | simpleCell : ∀{i} (s : String) → CellC i 64 | force (method (simpleCell s) get) = 65 | do' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 66 | return (s , simpleCell s) 67 | force (method (simpleCell _) (put s)) = 68 | do' (putStrLn ("putting (" ++ s ++ ")")) λ _ → 69 | return (_ , simpleCell s) 70 | 71 | 72 | program : ∀{i} → CellC ∞ → IO ConsoleInterface i Unit 73 | force (program c) = 74 | do' getLine λ s → 75 | method c (put s) >>= λ{ (_ , c) → 76 | method c get >>= λ{ (s' , c) → 77 | do (putStrLn s') λ _ → 78 | program c }} 79 | 80 | 81 | main : NativeIO Unit 82 | main = translateIOConsole (program (simpleCell "Start")) 83 | 84 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/objectsInAgdaUnsized.agda: -------------------------------------------------------------------------------- 1 | module objectsInAgdaUnsized where 2 | 3 | open import Data.Product 4 | open import interactiveProgramsAgdaUnsized hiding (main) 5 | open import NativeIO 6 | open import Data.String.Base 7 | 8 | record Interface : Set₁ where 9 | field Method : Set 10 | Result : (m : Method) → Set 11 | 12 | open Interface public 13 | 14 | record Object (I : Interface) : Set where 15 | coinductive 16 | field objectMethod : (m : Method I) → Result I m × Object I 17 | 18 | open Object public 19 | 20 | record IOObject (Iᵢₒ : IOInterface) (I : Interface) : Set where 21 | coinductive 22 | field method : (m : Method I) → IO Iᵢₒ (Result I m × IOObject Iᵢₒ I) 23 | 24 | open IOObject public 25 | 26 | data CellMethod A : Set where 27 | get : CellMethod A 28 | put : A → CellMethod A 29 | 30 | CellResult : ∀{A} → CellMethod A → Set 31 | CellResult {A} get = A 32 | CellResult (put _) = Unit 33 | 34 | cellI : (A : Set) → Interface 35 | Method (cellI A) = CellMethod A 36 | Result (cellI A) m = CellResult m 37 | 38 | 39 | {- A cell as an Object -} 40 | 41 | Cell : Set → Set 42 | Cell A = Object (cellI A) 43 | 44 | cell : {A : Set} → A → Cell A 45 | objectMethod (cell a) get = a , cell a 46 | objectMethod (cell a) (put a') = unit , cell a' 47 | 48 | 49 | {- A cell as an IOObject -} 50 | 51 | CellC : Set 52 | CellC = IOObject ConsoleInterface (cellI String) 53 | 54 | 55 | simpleCell : (s : String) → CellC 56 | force (method (simpleCell s) get) = 57 | do' (putStrLn ("getting (" ++ s ++ ")")) λ _ → 58 | delay (return' (s , simpleCell s)) 59 | force (method (simpleCell s) (put x)) = 60 | do' (putStrLn ("putting (" ++ x ++ ")")) λ _ → 61 | delay (return' (_ , simpleCell x)) 62 | 63 | {-# TERMINATING #-} 64 | 65 | program : CellC → IOConsole Unit 66 | force (program c₁) = 67 | do' getLine λ s → 68 | method c₁ (put s) >>= λ{ (_ , c₂) → 69 | method c₂ get >>= λ{ (s' , c₃) → 70 | do (putStrLn s') λ _ → 71 | (program c₃)}} 72 | 73 | main : NativeIO Unit 74 | main = translateIOConsole (program (simpleCell "Start")) 75 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/ship.ico: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/ooAgda/7cc45e0148a4a508d20ed67e791544c30fecd795/presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/ship.ico -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/spaceShipAdvanced.agda: -------------------------------------------------------------------------------- 1 | module spaceShipAdvanced where 2 | 3 | open import Unit 4 | open import Data.List.Base hiding ([]) 5 | open import Data.Bool.Base 6 | open import Data.Integer 7 | open import Data.Product hiding (map) 8 | open import NativeIO 9 | open import Size 10 | 11 | open import SizedIO.Object 12 | open import SizedIO.IOObject 13 | open import StateSizedIO.GUI.WxBindingsFFI 14 | open import StateSizedIO.GUI.WxGraphicsLib 15 | open import StateSized.GUI.ShipBitMap 16 | open import SizedIO.Base 17 | open import StateSizedIO.GUI.BaseStateDependent 18 | open import StateSizedIO.GUI.VariableList 19 | open import Sized.SimpleCell hiding (main; program) 20 | 21 | data GraphicServerMethod : Set where 22 | onPaintM : DC → Rect → GraphicServerMethod 23 | moveSpaceShipM : Frame → GraphicServerMethod 24 | callRepaintM : Frame → GraphicServerMethod 25 | 26 | GraphicServerResult : GraphicServerMethod → Set 27 | GraphicServerResult _ = Unit 28 | 29 | GraphicServerInterface : Interface 30 | Method GraphicServerInterface = GraphicServerMethod 31 | Result GraphicServerInterface = GraphicServerResult 32 | 33 | GraphicServerObject : ∀{i} → Set 34 | GraphicServerObject {i} = IOObject GuiLev1Interface GraphicServerInterface i 35 | 36 | graphicServerObject : ∀{i} → ℤ → GraphicServerObject {i} 37 | method (graphicServerObject z) (onPaintM dc rect) = 38 | do (drawBitmap dc ship (z , (+ 150)) true) λ _ → 39 | return (_ , graphicServerObject z) 40 | method (graphicServerObject z) (moveSpaceShipM fra) = 41 | return (_ , graphicServerObject (z + + 20)) 42 | method (graphicServerObject z) (callRepaintM fra) = 43 | do (repaint fra) λ _ → 44 | return (_ , graphicServerObject z) 45 | 46 | VarType = GraphicServerObject {∞} 47 | 48 | varInit : VarType 49 | varInit = graphicServerObject (+ 150) 50 | 51 | 52 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 53 | onPaint obj dc rect = mapIO proj₂ (method obj (onPaintM dc rect)) 54 | 55 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 56 | moveSpaceShip fra obj = mapIO proj₂ (method obj (moveSpaceShipM fra)) 57 | 58 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 59 | callRepaint fra obj = mapIO proj₂ (method obj (callRepaintM fra)) 60 | 61 | 62 | buttonHandler : ∀{i} → Frame → List (VarType → IO GuiLev1Interface i VarType) 63 | buttonHandler fra = moveSpaceShip fra ∷ [ callRepaint fra ] 64 | 65 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 66 | program = doˢ (level1C makeFrame) λ fra → 67 | doˢ (level1C (makeButton fra)) λ bt → 68 | doˢ (level1C (addButton fra bt)) λ _ → 69 | doˢ (createVar varInit) λ _ → 70 | doˢ (setButtonHandler bt (buttonHandler fra)) λ _ → 71 | doˢ (setOnPaint fra [ onPaint ]) 72 | returnˢ 73 | 74 | main : NativeIO Unit 75 | main = start (translateLev2 program) 76 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/spaceShipCell.agda: -------------------------------------------------------------------------------- 1 | module spaceShipCell where 2 | 3 | open import Unit 4 | open import Data.List.Base hiding ([]) 5 | open import Data.Bool.Base 6 | open import Data.Integer 7 | open import Data.Product hiding (map) 8 | open import NativeIO 9 | 10 | open import SizedIO.Object 11 | open import SizedIO.IOObject 12 | open import StateSizedIO.GUI.WxBindingsFFI 13 | open import StateSizedIO.GUI.WxGraphicsLib 14 | open import StateSized.GUI.ShipBitMap 15 | open import SizedIO.Base 16 | open import StateSizedIO.GUI.BaseStateDependent 17 | open import StateSizedIO.GUI.VariableList 18 | open import Sized.SimpleCell hiding (main; program) 19 | 20 | VarType = Object (cellI ℤ) 21 | 22 | cellℤC : (z : ℤ ) → VarType 23 | objectMethod (cellℤC z) get = ( z , cellℤC z ) 24 | objectMethod (cellℤC z) (put z') = ( _ , cellℤC z' ) 25 | 26 | varInit : VarType 27 | varInit = cellℤC (+ 150) 28 | 29 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 30 | onPaint c dc rect = 31 | let (z , c₁) = objectMethod c get in 32 | do (drawBitmap dc ship (z , (+ 150)) true) λ _ → 33 | return c₁ 34 | 35 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 36 | moveSpaceShip fra c = 37 | let (z , c₁) = objectMethod c get 38 | (_ , c₂) = objectMethod c₁ (put (z + + 20)) 39 | in return c₂ 40 | 41 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 42 | callRepaint fra c = do (repaint fra) λ _ → return c 43 | 44 | 45 | buttonHandler : ∀{i} → Frame → List (VarType → IO GuiLev1Interface i VarType) 46 | buttonHandler fra = moveSpaceShip fra ∷ [ callRepaint fra ] 47 | 48 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 49 | program = doˢ (level1C makeFrame) λ fra → 50 | doˢ (level1C (makeButton fra)) λ bt → 51 | doˢ (level1C (addButton fra bt)) λ _ → 52 | doˢ (createVar varInit) λ _ → 53 | doˢ (setButtonHandler bt (buttonHandler fra)) λ _ → 54 | doˢ (setOnPaint fra [ onPaint ]) 55 | returnˢ 56 | 57 | main : NativeIO Unit 58 | main = start (translateLev2 program) 59 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/spaceShipSimpleVar.agda: -------------------------------------------------------------------------------- 1 | module spaceShipSimpleVar where 2 | 3 | open import Unit 4 | open import Data.List.Base hiding ([]) 5 | open import Data.Bool.Base 6 | open import Data.Integer 7 | open import Data.Product hiding (map) 8 | open import NativeIO 9 | 10 | open import StateSizedIO.GUI.WxBindingsFFI 11 | open import StateSizedIO.GUI.WxGraphicsLib 12 | open import StateSized.GUI.ShipBitMap 13 | open import SizedIO.Base 14 | open import StateSizedIO.GUI.BaseStateDependent 15 | open import StateSizedIO.GUI.VariableList 16 | 17 | VarType = ℤ 18 | 19 | varInit : VarType 20 | varInit = (+ 150) 21 | 22 | onPaint : ∀{i} → VarType → DC → Rect → IO GuiLev1Interface i VarType 23 | onPaint z dc rect = do (drawBitmap dc ship (z , (+ 150)) true) λ _ → 24 | return z 25 | 26 | moveSpaceShip : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 27 | moveSpaceShip fra z = return (z + + 20) 28 | 29 | callRepaint : ∀{i} → Frame → VarType → IO GuiLev1Interface i VarType 30 | callRepaint fra z = do (repaint fra) λ _ → return z 31 | 32 | buttonHandler : ∀{i} → Frame → List (VarType → IO GuiLev1Interface i VarType) 33 | buttonHandler fra = moveSpaceShip fra ∷ [ callRepaint fra ] 34 | 35 | program : ∀{i} → IOˢ GuiLev2Interface i (λ _ → Unit) [] 36 | program = doˢ (level1C makeFrame) λ fra → 37 | doˢ (level1C (makeButton fra)) λ bt → 38 | doˢ (level1C (addButton fra bt)) λ _ → 39 | doˢ (createVar varInit) λ _ → 40 | doˢ (setButtonHandler bt (buttonHandler fra)) λ _ → 41 | doˢ (setOnPaint fra [ onPaint ]) 42 | returnˢ 43 | 44 | main : NativeIO Unit 45 | main = start (translateLev2 program) 46 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/stateDependentIO.agda: -------------------------------------------------------------------------------- 1 | module stateDependentIO where 2 | 3 | open import Level using (_⊔_ ) renaming (suc to lsuc; zero to lzero) 4 | open import Size renaming (Size to AgdaSize) 5 | open import NativeIO 6 | open import Function 7 | 8 | module _ {γ ρ μ} where 9 | record IOInterfaceˢ : Set (lsuc (γ ⊔ ρ ⊔ μ )) where 10 | field 11 | StateIOˢ : Set γ 12 | Commandˢ : StateIOˢ → Set ρ 13 | Responseˢ : (s : StateIOˢ) → Commandˢ s → Set μ 14 | nextIOˢ : (s : StateIOˢ) → (c : Commandˢ s) → Responseˢ s c 15 | → StateIOˢ 16 | 17 | open IOInterfaceˢ public 18 | 19 | module _ {α γ ρ μ}(i : IOInterfaceˢ {γ} {ρ} {μ} ) 20 | (let S = StateIOˢ i) (let C = Commandˢ i) 21 | (let R = Responseˢ i) (let next = nextIOˢ i) 22 | where 23 | mutual 24 | 25 | record IOˢ (i : AgdaSize) (A : S → Set α) (s : S) 26 | : Set (lsuc (α ⊔ γ ⊔ ρ ⊔ μ )) where 27 | coinductive 28 | constructor delay 29 | field forceˢ : {j : Size< i} → IOˢ' j A s 30 | 31 | data IOˢ' (i : AgdaSize) (A : S → Set α) : S 32 | → Set (lsuc (α ⊔ γ ⊔ ρ ⊔ μ )) where 33 | doˢ' : {s : S} → (c : C s) → (f : (r : R s c) 34 | → IOˢ i A (next s c r) ) 35 | → IOˢ' i A s 36 | returnˢ' : {s : S} → (a : A s) → IOˢ' i A s 37 | 38 | 39 | open IOˢ public 40 | 41 | module _ {α γ ρ μ}{I : IOInterfaceˢ {γ} {ρ} {μ}} 42 | (let S = StateIOˢ I) (let C = Commandˢ I) 43 | (let R = Responseˢ I) (let next = nextIOˢ I) where 44 | 45 | 46 | returnIOˢ : ∀{i}{A : S → Set α} {s : S} (a : A s) → IOˢ I i A s 47 | forceˢ (returnIOˢ a) = returnˢ' a 48 | 49 | doˢ : ∀{i}{A : S → Set α} {s : S} 50 | (c : C s) (f : (r : R s c) → IOˢ I i A (next s c r)) → IOˢ I i A s 51 | forceˢ (doˢ c f) = doˢ' c f 52 | 53 | 54 | 55 | module _ {γ ρ}{I : IOInterfaceˢ {γ} {ρ} {lzero}} 56 | (let S = StateIOˢ I) (let C = Commandˢ I) 57 | (let R = Responseˢ I) (let next = nextIOˢ I) where 58 | 59 | {-# NON_TERMINATING #-} 60 | translateIOˢ : ∀{A : Set }{s : S} 61 | → (translateLocal : (s : S) → (c : C s) → NativeIO (R s c)) 62 | → IOˢ I ∞ (λ s → A) s 63 | → NativeIO A 64 | translateIOˢ {A} {s} translateLocal p = case (forceˢ p {_}) of 65 | λ{ (doˢ' {.s} c f) → (translateLocal s c) native>>= λ r → 66 | translateIOˢ translateLocal (f r) 67 | ; (returnˢ' a) → nativeReturn a 68 | } 69 | -------------------------------------------------------------------------------- /presentationsAndExampleCode/agdaImplementorsMeetingGlasgow22April2016AntonSetzer/talkGlasgow.agda-lib: -------------------------------------------------------------------------------- 1 | name: talkGlasgow 2 | depend: std-lib ooAgda ooAgdaExamples 3 | include: . 4 | -------------------------------------------------------------------------------- /src/ConsoleLib.agda: -------------------------------------------------------------------------------- 1 | module ConsoleLib where 2 | 3 | open import NativeIO public 4 | open import SizedIO.Console public hiding (main) renaming (translateIOConsole to run) 5 | open import Size 6 | open import SizedIO.Base 7 | open import Data.List 8 | 9 | WriteString : (s : String) → IOConsole ∞ Unit 10 | WriteString s = Exec (putStrLn s) 11 | 12 | GetLine : IOConsole ∞ String 13 | GetLine = Exec getLine 14 | 15 | ConsoleProg : Set 16 | ConsoleProg = NativeIO Unit 17 | -------------------------------------------------------------------------------- /src/NativeIO.agda: -------------------------------------------------------------------------------- 1 | module NativeIO where 2 | 3 | open import Unit public 4 | open import Data.List 5 | open import Data.Nat 6 | open import Data.String.Base using (String) public 7 | 8 | {-# FOREIGN GHC import qualified Data.Text #-} 9 | {-# FOREIGN GHC import qualified System.Environment #-} 10 | 11 | 12 | postulate 13 | NativeIO : Set → Set 14 | nativeReturn : {A : Set} → A → NativeIO A 15 | _native>>=_ : {A B : Set} → NativeIO A → (A → NativeIO B) → NativeIO B 16 | _native>>_ : {A B : Set} → NativeIO A → NativeIO B → NativeIO B 17 | 18 | 19 | {-# BUILTIN IO NativeIO #-} 20 | {-# COMPILE GHC NativeIO = type IO #-} 21 | {-# COMPILE GHC nativeReturn = (\_ -> return :: a -> IO a) #-} 22 | {-# COMPILE GHC _native>>=_ = (\_ _ -> (>>=) :: IO a -> (a -> IO b) -> IO b) #-} 23 | {-# COMPILE GHC _native>>_ = (\_ _ -> (>>) :: IO a -> IO b -> IO b) #-} 24 | 25 | 26 | 27 | 28 | postulate 29 | nativeGetLine : NativeIO String 30 | nativePutStrLn : String → NativeIO Unit 31 | 32 | {-# COMPILE GHC nativePutStrLn = (\ s -> putStrLn (Data.Text.unpack s)) #-} 33 | {-# COMPILE GHC nativeGetLine = (fmap Data.Text.pack getLine) #-} 34 | 35 | 36 | postulate 37 | getArgs : NativeIO (List String) 38 | primShowNat : ℕ → String 39 | 40 | {-# COMPILE GHC primShowNat = Data.Text.pack . show #-} 41 | 42 | {-# COMPILE GHC getArgs = fmap (map Data.Text.pack) System.Environment.getArgs #-} 43 | 44 | 45 | 46 | -- 47 | -- Debug function 48 | -- 49 | {-# FOREIGN GHC import qualified Debug.Trace as Trace #-} 50 | 51 | {-# FOREIGN GHC 52 | 53 | debug = flip Trace.trace 54 | 55 | debug' :: c -> Data.Text.Text -> c 56 | debug' c txt = debug c (Data.Text.unpack txt) 57 | 58 | #-} 59 | 60 | postulate debug : {A : Set} → A → String → A 61 | --debug {A} a s = a 62 | {-# COMPILE GHC debug = (\x -> debug') #-} 63 | 64 | postulate debug₁ : {A : Set₁} → A → String → A 65 | -- debug₁ {A} a s = a 66 | {-# COMPILE GHC debug₁ = (\x -> debug') #-} 67 | -------------------------------------------------------------------------------- /src/NativeInt.agda: -------------------------------------------------------------------------------- 1 | module NativeInt where 2 | 3 | open import Data.Integer.Base hiding (_+_) 4 | 5 | postulate 6 | Int : Set 7 | toInt : ℤ -> Int 8 | fromInt : Int -> ℤ 9 | 10 | {-# COMPILE GHC Int = type Int #-} 11 | {-# COMPILE GHC toInt fromInteger #-} 12 | {-# COMPILE GHC fromInt toInteger #-} 13 | -------------------------------------------------------------------------------- /src/NativePolyIO.agda: -------------------------------------------------------------------------------- 1 | module NativePolyIO where 2 | 3 | open import Data.String.Base using (String) public 4 | open import Level 5 | 6 | 7 | record Unit {α} : Set α where 8 | constructor unit 9 | 10 | {-# HASKELL type AgdaUnit a = () #-} 11 | 12 | {-# COMPILED_DATA Unit AgdaUnit () #-} 13 | 14 | postulate 15 | NativeIO : ∀ {ℓ} → Set ℓ → Set ℓ 16 | nativeReturn : ∀ {a} {A : Set a} → A → NativeIO A 17 | _native>>=_ : ∀ {a b} {A : Set a} {B : Set b} → NativeIO A → (A → NativeIO B) → NativeIO B 18 | 19 | 20 | {-# COMPILED_TYPE NativeIO MAlonzo.Code.NativePolyIO.AgdaIO #-} 21 | {-# BUILTIN IO NativeIO #-} 22 | 23 | {-# HASKELL type AgdaIO a b = IO b #-} 24 | 25 | 26 | {-# COMPILED nativeReturn (\_ _ -> return :: a -> IO a) #-} 27 | {-# COMPILED _native>>=_ (\_ _ _ _ -> 28 | (>>=) :: IO a -> (a -> IO b) -> IO b) #-} 29 | 30 | {-# IMPORT Data.Text.IO #-} 31 | 32 | postulate 33 | nativePutStrLn : ∀ {ℓ} → String → NativeIO (Unit{ℓ}) 34 | nativeGetLine : NativeIO String 35 | 36 | 37 | 38 | 39 | {-# COMPILED nativePutStrLn (\ _ s -> Data.Text.IO.putStrLn s) #-} 40 | {-# COMPILED nativeGetLine Data.Text.IO.getLine #-} 41 | -------------------------------------------------------------------------------- /src/SizedIO/Console.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.Console where 2 | 3 | open import Level using () renaming (zero to lzero) 4 | open import Size 5 | 6 | open import NativeIO 7 | open import SizedIO.Base 8 | 9 | data ConsoleCommand : Set where 10 | putStrLn : String → ConsoleCommand 11 | getLine : ConsoleCommand 12 | 13 | ConsoleResponse : ConsoleCommand → Set 14 | ConsoleResponse (putStrLn s) = Unit 15 | ConsoleResponse getLine = String 16 | 17 | consoleI : IOInterface 18 | Command consoleI = ConsoleCommand 19 | Response consoleI = ConsoleResponse 20 | 21 | IOConsole : Size → Set → Set 22 | IOConsole i = IO consoleI i 23 | 24 | IOConsole+ : Size → Set → Set 25 | IOConsole+ i = IO+ consoleI i 26 | 27 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 28 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 29 | translateIOConsoleLocal getLine = nativeGetLine 30 | 31 | translateIOConsole : {A : Set} → IOConsole ∞ A → NativeIO A 32 | translateIOConsole = translateIO translateIOConsoleLocal 33 | 34 | postulate translateIOConsoleTODO : {A : Set} → (i : Size) → IOConsole i A → NativeIO A 35 | 36 | 37 | main : NativeIO Unit 38 | main = nativePutStrLn "Console" 39 | -------------------------------------------------------------------------------- /src/SizedIO/ConsoleObject.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.ConsoleObject where 2 | 3 | open import Size 4 | 5 | open import SizedIO.Console 6 | open import SizedIO.Object 7 | open import SizedIO.IOObject 8 | 9 | 10 | -- A console object is an IO object for the IO interface of console 11 | ConsoleObject : (i : Size) → (iface : Interface) → Set 12 | ConsoleObject i iface = IOObject consoleI iface i 13 | 14 | -------------------------------------------------------------------------------- /src/SizedIO/IOObject.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.IOObject where 2 | 3 | open import Data.Product 4 | 5 | open import Size 6 | 7 | open import SizedIO.Base 8 | open import SizedIO.Object 9 | 10 | -- An IO object is like a simple object, 11 | -- but the method returns IO applied to the result type of a simple object 12 | -- which means the method returns an IO program which when terminating 13 | -- returns the result of the simple object 14 | 15 | module _ (ioi : IOInterface) (let C = Command ioi) (let R = Response ioi) 16 | (oi : Interface) (let M = Method oi) (let Rt = Result oi) 17 | where 18 | 19 | record IOObject (i : Size) : Set where 20 | coinductive 21 | field 22 | method : ∀{j : Size< i} (m : M) → IO ioi ∞ (Rt m × IOObject j) 23 | 24 | open IOObject public 25 | -------------------------------------------------------------------------------- /src/SizedIO/Object.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.Object where 2 | 3 | open import Data.Product 4 | 5 | record Interface : Set₁ where 6 | field 7 | Method : Set 8 | Result : (m : Method) → Set 9 | open Interface public 10 | 11 | -- A simple object just returns for a method the response 12 | -- and the object itself 13 | record Object (i : Interface) : Set where 14 | coinductive 15 | field 16 | objectMethod : (m : Method i) → Result i m × Object i 17 | 18 | open Object public 19 | 20 | 21 | _▹_ : {A : Set} → {B : Set} → A → (A → B) → B 22 | a ▹ f = f a 23 | -------------------------------------------------------------------------------- /src/SizedIO/coIOIO.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.coIOIO where 2 | 3 | open import Size 4 | 5 | 6 | mutual 7 | data coIO² (i : Size) (j : Size) 8 | (Cin : Set ) (Rin : Cin → Set) 9 | (Cext : Set) (Rext : Cext → Set) 10 | (A : Set) : Set where 11 | return : A → coIO² i j Cin Rin Cext Rext A 12 | dof : (i' : Size< i) → 13 | (c : Cin) → (Rin c → coIO² i' j Cin Rin Cext Rext A) 14 | → coIO² i j Cin Rin Cext Rext A 15 | do∞ : (c : Cext) 16 | → (Rext c → coIO²∞ j Cin Rin Cext Rext A) 17 | → coIO² i j Cin Rin Cext Rext A 18 | 19 | record coIO²∞ (j : Size) (Cin : Set ) (Rin : Cin → Set) 20 | (Cext : Set) (Rext : Cext → Set) 21 | (A : Set) : Set where 22 | coinductive 23 | field 24 | force : (j' : Size< j) 25 | → coIO² ∞ j' Cin Rin Cext Rext A 26 | 27 | open coIO²∞ public 28 | -------------------------------------------------------------------------------- /src/SizedIO/coIOIOObject.agda: -------------------------------------------------------------------------------- 1 | module SizedIO.coIOIOObject where 2 | 3 | open import Data.Unit.Base 4 | open import Data.Product 5 | open import Data.String.Base 6 | 7 | open import Size 8 | 9 | open import SizedIO.Object 10 | 11 | open import SizedIO.Base 12 | open import SizedIO.Console 13 | open import SizedIO.coIOIO 14 | -- open import SizedIO.IOObject 15 | 16 | 17 | record IOObject' (i : Size) (iface : Interface)(C : Set)(R : C → Set) : Set 18 | where 19 | coinductive 20 | field 21 | method : ∀{j : Size< i} (m : Method iface) → 22 | IO j C R (Response iface m × IOObject' j iface C R) 23 | open IOObject' public 24 | 25 | 26 | 27 | -- An IO object is like a simple object, 28 | -- but the method returns IO applied to the result type of a simple object 29 | -- which means the method returns an IO program which when terminating 30 | -- returns the result of the simple object 31 | record coIO²Object (i : Size) (iface : Interface) 32 | (Cin : Set)(Rin : Cin → Set) 33 | (Cext : Set)(Rext : Cext → Set) : Set 34 | where 35 | coinductive 36 | field 37 | method' : ∀{j : Size< i} (m : Method iface) → 38 | coIO²∞ ∞ Cin Rin Cext Rext 39 | (Response iface m × coIO²Object j iface Cin Rin Cext Rext) 40 | open coIO²Object public 41 | 42 | 43 | selfrefIOObject : (i : Size) (iface : Interface) 44 | (Cext : Set) (Rext : Cext → Set) → Set 45 | selfrefIOObject i iface Cext Rext = coIO²Object i iface (Method iface) 46 | (Response iface) Cext Rext 47 | 48 | mutual 49 | compileSelfRef : (i : Size) (iface : Interface) 50 | (Cext : Set) (Rext : Cext → Set) 51 | (obj : selfrefIOObject i iface Cext Rext) → 52 | IOObject' i iface Cext Rext 53 | IO.force (method (compileSelfRef i iface Cext Rext obj) {i'} m) {j} = 54 | compileSelfRefaux i' j iface Cext Rext (Response iface m) (coIO²∞.force (method' obj {i'} m) j) 55 | 56 | compileSelfRefaux : 57 | (i : Size) (j : Size< i) 58 | (iface : Interface) (let M = Method iface) (let R = Response iface) 59 | (Cext : Set) (Rext : Cext → Set) 60 | (A : Set) 61 | (coObj : coIO² ∞ j M R Cext Rext (A × coIO²Object i iface M R Cext Rext)) 62 | → IO' j Cext Rext (A × IOObject' i iface Cext Rext) 63 | compileSelfRefaux i j iface Cext Rext A (coIO².return (r , obj)) = return' (r , compileSelfRef i iface Cext Rext obj) 64 | compileSelfRefaux i j iface Cext Rext A (dof i' m f) = {! >>=!} 65 | compileSelfRefaux i j iface Cext Rext A (do∞ c f) = 66 | do' c (λ r → compileSelfRefaux'' i j iface Cext Rext A (f r)) 67 | 68 | 69 | 70 | compileSelfRefaux'' : (i : Size) (j : Size< i) 71 | (iface : Interface) (let M = Method iface) (let R = Response iface) 72 | (Cext : Set) (Rext : Cext → Set) 73 | (A : Set) 74 | (coObj : coIO²∞ j M R Cext Rext (A × coIO²Object i iface M R Cext Rext)) 75 | → IO j Cext Rext (A × IOObject' i iface Cext Rext) 76 | IO.force (compileSelfRefaux'' i j iface Cext Rext A coObj) {j''} = compileSelfRefaux i j'' iface Cext Rext A (coIO²∞.force coObj j'') 77 | 78 | 79 | -- compileSelfRefaux'' i' j iface Cext Rext A coObj 80 | -- = compileSelfRefaux' i' j iface Cext Rext A (coIO²∞.force coObj {!!}) 81 | 82 | {- 83 | compileSelfRefaux' : 84 | (i' : Size) (j : Size< i') 85 | (iface : Interface) (let M = Methods iface) (let R = Responses iface) 86 | (Cext : Set) (Rext : Cext → Set) 87 | (A : Set) 88 | (coObj : coIO² ∞ j M R Cext Rext (A × coIO²Object i' iface M R Cext Rext)) 89 | → IO j Cext Rext (A × IOObject' i' iface Cext Rext) 90 | IO.force (compileSelfRefaux' i' j iface Cext Rext A coObj) = compileSelfRefaux i' j iface Cext Rext A coObj 91 | -} 92 | {- 93 | 123;456hallo789 94 | 123;456hallo789 95 | 123;789hallo456 96 | 123;789hallo456 97 | 123;789456 98 | 123;789456 99 | 123;456789 100 | 123;456789 101 | 123;456789 102 | 123;456789 103 | 104 | C-x r k -- kill rectangle 105 | C-x r y -- yank rectangle 106 | C-x r t -- text rectangle 107 | -} 108 | 109 | {- 110 | M-x abbrev-mode 111 | 112 | C-x ail 113 | 114 | crosss 115 | → 116 | 117 | \ 118 | -} 119 | 120 | 121 | -- method' obj {j'} m 122 | 123 | {- 124 | compileSelfRefaux : (j' : Size) (iface : Interface) 125 | (Cext : Set) (Rext : Cext → Set) 126 | (m : Methods iface) 127 | ( p : coIO²∞ ∞ (Methods iface) 128 | (Responses iface) Cext Rext 129 | (Responses iface m × 130 | coIO²Object j' iface (Methods iface) 131 | (Responses iface) Cext Rext)) 132 | → IO' j' Cext Rext 133 | (Responses iface m × 134 | IOObject' j iface Cext Rext)IO' j' Cext Rext 135 | (Responses iface m × IOObject' j iface Cext Rext) 136 | 137 | -} 138 | -------------------------------------------------------------------------------- /src/SizedPolyIO/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --copatterns #-} 2 | 3 | module SizedPolyIO.Base where 4 | 5 | open import Data.Maybe.Base 6 | open import Data.Sum renaming (inj₁ to left; inj₂ to right; [_,_]′ to either) 7 | 8 | open import Function 9 | open import Level using (_⊔_) renaming (suc to lsuc) 10 | open import Size 11 | 12 | open import NativePolyIO 13 | 14 | record IOInterface γ ρ : Set (lsuc (γ ⊔ ρ)) where 15 | field 16 | Command : Set γ 17 | Response : (m : Command) → Set ρ 18 | open IOInterface public 19 | 20 | module _ {γ ρ} (I : IOInterface γ ρ) (let C = Command I) (let R = Response I) where 21 | 22 | mutual 23 | record IO (i : Size) {α} (A : Set α) : Set (α ⊔ γ ⊔ ρ) where 24 | coinductive 25 | constructor delay 26 | field 27 | force : {j : Size< i} → IO' j A 28 | 29 | data IO' (i : Size) {α} (A : Set α) : Set (α ⊔ γ ⊔ ρ) where 30 | do' : (c : C) (f : R c → IO i A) → IO' i A 31 | return' : (a : A) → IO' i A 32 | 33 | data IO+ (i : Size) {α} (A : Set α) : Set (α ⊔ ρ ⊔ γ) where 34 | do' : (c : C) (f : R c → IO i A) → IO+ i A 35 | 36 | open IO public 37 | 38 | module _ {γ ρ} {I : IOInterface γ ρ} (let C = Command I) (let R = Response I) where 39 | 40 | return : ∀{i α}{A : Set α} (a : A) → IO I i A 41 | force (return a) = return' a 42 | 43 | do : ∀{i α}{A : Set α} (c : C) (f : R c → IO I i A) → IO I i A 44 | force (do c f) = do' c f 45 | 46 | do1 : ∀{i} (c : C) → IO I i (R c) 47 | do1 c = do c return 48 | 49 | infixl 2 _>>=_ _>>='_ _>>_ 50 | 51 | mutual 52 | _>>='_ : ∀{i α}{A B : Set α} (m : IO' I i A) (k : A → IO I (↑ i) B) → IO' I i B 53 | do' c f >>=' k = do' c λ x → f x >>= k 54 | return' a >>=' k = force (k a) 55 | 56 | _>>=_ : ∀{i α}{A B : Set α} (m : IO I i A) (k : A → IO I i B) → IO I i B 57 | force (m >>= k) = force m >>=' k 58 | 59 | _>>_ : ∀{i α}{B : Set α} (m : IO I i Unit) (k : IO I i B) → IO I i B 60 | m >> k = m >>= λ _ → k 61 | 62 | {-# NON_TERMINATING #-} 63 | translateIO : ∀{α}{A : Set α} 64 | → (translateLocal : (c : C) → NativeIO (R c)) 65 | → IO I ∞ A 66 | → NativeIO A 67 | translateIO translateLocal m = case (force m {_}) of 68 | λ{ (do' c f) → (translateLocal c) native>>= λ r → 69 | translateIO translateLocal (f r) 70 | ; (return' a) → nativeReturn a 71 | } 72 | 73 | -- Recursion 74 | 75 | -- trampoline provides a generic form of loop (generalizing while/repeat). 76 | -- Starting at state s : S, step function f is iterated until it returns 77 | -- a result in A. 78 | 79 | fromIO+' : ∀{i α}{A : Set α} → IO+ I i A → IO' I i A 80 | fromIO+' (do' c f) = do' c f 81 | 82 | fromIO+ : ∀{i α}{A : Set α} → IO+ I i A → IO I i A 83 | force (fromIO+ (do' c f)) = do' c f 84 | 85 | _>>=+'_ : ∀{i α}{A B : Set α} (m : IO+ I i A) (k : A → IO I i B) → IO' I i B 86 | do' c f >>=+' k = do' c λ x → f x >>= k 87 | 88 | _>>=+_ : ∀{i α}{A B : Set α} (m : IO+ I i A) (k : A → IO I i B) → IO I i B 89 | force (m >>=+ k) = m >>=+' k 90 | 91 | mutual 92 | 93 | _>>+_ : ∀{i α}{A B : Set α} (m : IO I i (A ⊎ B)) (k : A → IO I i B) → IO I i B 94 | force (m >>+ k) = force m >>+' k 95 | 96 | _>>+'_ : ∀{j α}{A B : Set α} (m : IO' I j (A ⊎ B)) (k : A → IO I (↑ j) B) → IO' I j B 97 | do' c f >>+' k = do' c λ x → f x >>+ k 98 | return' (left a) >>+' k = force (k a) 99 | return' (right b) >>+' k = return' b 100 | 101 | -- loop 102 | 103 | trampoline : ∀{i α}{A S : Set α} (f : S → IO+ I i (S ⊎ A)) (s : S) → IO I i A 104 | force (trampoline f s) = case (f s) of 105 | \{ (do' c k) → do' c λ r → k r >>+ trampoline f } 106 | 107 | -- simple infinite loop 108 | 109 | forever : ∀{i α}{A B : Set α} → IO+ I i A → IO I i B 110 | force (forever (do' c f)) = do' c λ r → f r >>= λ _ → forever (do' c f) 111 | 112 | whenJust : ∀{i α}{A : Set α} → Maybe A → (A → IO I i (Unit {α})) → IO I i Unit 113 | whenJust nothing k = return _ 114 | whenJust (just a) k = k a 115 | -------------------------------------------------------------------------------- /src/SizedPolyIO/Console.agda: -------------------------------------------------------------------------------- 1 | module SizedPolyIO.Console where 2 | 3 | open import Level using () renaming (zero to lzero) 4 | open import Size 5 | 6 | open import NativePolyIO 7 | open import SizedPolyIO.Base 8 | 9 | data ConsoleCommand : Set where 10 | putStrLn : String → ConsoleCommand 11 | getLine : ConsoleCommand 12 | 13 | ConsoleResponse : ConsoleCommand → Set 14 | ConsoleResponse (putStrLn s) = Unit 15 | ConsoleResponse getLine = String 16 | 17 | consoleI : IOInterface lzero lzero 18 | Command consoleI = ConsoleCommand 19 | Response consoleI = ConsoleResponse 20 | 21 | IOConsole : Size → Set → Set 22 | IOConsole i = IO consoleI i 23 | 24 | IOConsole+ : Size → Set → Set 25 | IOConsole+ i = IO+ consoleI i 26 | 27 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 28 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 29 | translateIOConsoleLocal getLine = nativeGetLine 30 | 31 | translateIOConsole : {A : Set} → IOConsole ∞ A → NativeIO A 32 | translateIOConsole = translateIO translateIOConsoleLocal 33 | 34 | main : NativeIO (Unit {lzero}) 35 | main = nativePutStrLn "Console" 36 | -------------------------------------------------------------------------------- /src/SizedPolyIO/ConsoleObject.agda: -------------------------------------------------------------------------------- 1 | module SizedPolyIO.ConsoleObject where 2 | 3 | open import Size 4 | open import Level using (_⊔_) renaming (suc to lsuc) 5 | 6 | open import SizedPolyIO.Console 7 | open import SizedPolyIO.Object 8 | open import SizedPolyIO.IOObject 9 | 10 | 11 | -- A console object is an IO object for the IO interface of console 12 | ConsoleObject : ∀{μ ρ}(i : Size) → (iface : Interface μ ρ) → Set (μ ⊔ ρ) 13 | ConsoleObject i iface = IOObject consoleI iface i 14 | 15 | -------------------------------------------------------------------------------- /src/SizedPolyIO/IOObject.agda: -------------------------------------------------------------------------------- 1 | module SizedPolyIO.IOObject where 2 | 3 | open import Data.Product 4 | 5 | open import Level using (_⊔_) renaming (suc to lsuc) 6 | open import Size 7 | 8 | open import SizedPolyIO.Base 9 | open import SizedPolyIO.Object 10 | 11 | -- An IO object is like a simple object, 12 | -- but the method returns IO applied to the result type of a simple object 13 | -- which means the method returns an IO program which when terminating 14 | -- returns the result of the simple object 15 | 16 | 17 | module _ {γ ρ} (ioi : IOInterface γ ρ) (let C = Command ioi) (let R = Response ioi) 18 | {μ σ} (oi : Interface μ σ) (let M = Method oi) (let Rt = Result oi) 19 | where 20 | 21 | record IOObject (i : Size) : Set (ρ ⊔ γ ⊔ μ ⊔ σ) where 22 | coinductive 23 | field 24 | method : ∀{j : Size< i} (m : M) → IO ioi ∞ (Rt m × IOObject j) 25 | 26 | open IOObject public 27 | -------------------------------------------------------------------------------- /src/SizedPolyIO/Object.agda: -------------------------------------------------------------------------------- 1 | module SizedPolyIO.Object where 2 | 3 | open import Data.Product 4 | 5 | open import Level using (_⊔_) renaming (suc to lsuc) 6 | 7 | record Interface μ ρ : Set (lsuc (μ ⊔ ρ)) where 8 | field 9 | Method : Set μ 10 | Result : (m : Method) → Set ρ 11 | open Interface public 12 | 13 | -- A simple object just returns for a method the response 14 | -- and the object itself 15 | record Object {μ ρ} (i : Interface μ ρ) : Set (μ ⊔ ρ) where 16 | coinductive 17 | field 18 | objectMethod : (m : Method i) → Result i m × Object i 19 | 20 | open Object public 21 | -------------------------------------------------------------------------------- /src/StateSizedIO/Base.agda: -------------------------------------------------------------------------------- 1 | 2 | module StateSizedIO.Base where 3 | 4 | open import Size 5 | open import SizedIO.Base 6 | 7 | open import Data.Product 8 | 9 | record IOInterfaceˢ : Set₁ where 10 | field 11 | IOStateˢ : Set 12 | Commandˢ : IOStateˢ → Set 13 | Responseˢ : (s : IOStateˢ) → (m : Commandˢ s) → Set 14 | IOnextˢ : (s : IOStateˢ) → (m : Commandˢ s) → (Responseˢ s m) → IOStateˢ 15 | open IOInterfaceˢ public 16 | 17 | 18 | record Interfaceˢ : Set₁ where 19 | field 20 | Stateˢ : Set 21 | Methodˢ : Stateˢ → Set 22 | Resultˢ : (s : Stateˢ) → (m : Methodˢ s) → Set 23 | nextˢ : (s : Stateˢ) → (m : Methodˢ s) → (Resultˢ s m) → Stateˢ 24 | open Interfaceˢ public 25 | 26 | 27 | module _ (ioi : IOInterface) (let C = Command ioi) (let R = Response ioi) 28 | (oi : Interfaceˢ) (let S = Stateˢ oi) (let M = Methodˢ oi) (let Rt = Resultˢ oi) 29 | (let n = nextˢ oi) 30 | where 31 | 32 | record IOObjectˢ (i : Size) (s : Stateˢ) : Set where 33 | coinductive 34 | field 35 | method : ∀{j : Size< i} (m : M s) → IO ioi ∞ ( Σ[ r ∈ Rt s m ] IOObjectˢ j (n s m r)) 36 | 37 | open IOObjectˢ public 38 | 39 | 40 | 41 | module _ (I : IOInterfaceˢ ) 42 | (let S = IOStateˢ I) (let C = Commandˢ I) 43 | (let R = Responseˢ I) (let n = nextˢ I) 44 | where 45 | 46 | mutual 47 | record IOˢ (i : Size) (A : S → Set) (s : S) : Set where 48 | coinductive 49 | constructor delay 50 | field 51 | forceˢ : {j : Size< i} → IOˢ' j A s 52 | 53 | data IOˢ' (i : Size) (A : S → Set) (s : S) : Set where 54 | doˢ' : (c : C s) (f : (r : R s c) → IOˢ i A (n s c r)) → IOˢ' i A s 55 | returnˢ' : (a : A s) → IOˢ' i A s 56 | 57 | data IOˢ+ (i : Size) (A : S → Set) (s : S) : Set where 58 | do' : (c : C s) (f : (r : R s c) → IOˢ i A (n s c r)) → IOˢ+ i A s 59 | 60 | open IOˢ public 61 | 62 | module _ {I : IOInterfaceˢ } 63 | (let S = Stateˢ I) (let C = Commandˢ I) 64 | (let R = Responseˢ I) (let n = nextˢ I) 65 | where 66 | returnˢ : ∀{i}{A : S → Set} (s : S) (a : A s) → IOˢ I i A s 67 | forceˢ (returnˢ s a) = returnˢ' a 68 | 69 | doˢ : ∀{i}{A : S → Set} (s : S) (c : C s) 70 | (f : (r : R s c) → IOˢ I i A (n s c r)) 71 | → IOˢ I i A s 72 | forceˢ (doˢ s c f) = doˢ' c f 73 | -------------------------------------------------------------------------------- /src/StateSizedIO/Base.agda.README.txt: -------------------------------------------------------------------------------- 1 | 2 | "StateSizedIO.Base" 3 | 4 | Change: State, Methods, etc. all have different levels 5 | to get the old code running, I need to add those levels! 6 | 7 | Base should no longer be used 8 | there is a more level generic one in 9 | StateSizedIO.GUI.BaseStateDependent 10 | 11 | The old version of Base.agda is in 12 | BaseOld.agda 13 | -------------------------------------------------------------------------------- /src/StateSizedIO/GUI/VariableList.agda: -------------------------------------------------------------------------------- 1 | module StateSizedIO.GUI.VariableList where 2 | 3 | open import Data.Product hiding (map) 4 | open import Data.List 5 | 6 | open import NativeIO 7 | open import StateSizedIO.GUI.WxBindingsFFI using (Var; nativeTakeVar; nativePutVar) 8 | 9 | 10 | 11 | data VarList : Set₁ where 12 | [] : VarList 13 | addVar : (A : Set) → Var A → VarList → VarList 14 | 15 | addVar' : {A : Set} → Var A → VarList → VarList 16 | addVar' {A} = addVar A 17 | 18 | [_]var : ∀ {A : Set} → Var A → VarList 19 | [_]var{A} x = addVar A x [] 20 | 21 | 22 | prod : VarList → Set 23 | prod [] = Unit 24 | prod (addVar A v []) = A 25 | prod (addVar A v l) = A × prod l 26 | 27 | varListProj₁ : (l : VarList){A : Set}{v : Var A}(p : prod (addVar A v l)) → A 28 | varListProj₁ [] {A} {v} a = a 29 | varListProj₁ (addVar A₁ x l) {A} {v} (a , rest) = a 30 | 31 | 32 | varListUpdateProj₁ : (l : VarList){A : Set}{v : Var A}(a : A) 33 | (p : prod (addVar A v l)) 34 | → prod (addVar A v l) 35 | varListUpdateProj₁ [] {A} {v} a p = a 36 | varListUpdateProj₁ (addVar A₁ x l) {A} {v} a' (a , rest) = ( a' , rest ) 37 | 38 | 39 | takeVar : (l : VarList) → NativeIO (prod l) 40 | takeVar [] = nativeReturn _ 41 | takeVar (addVar A v []) = nativeTakeVar {A} v native>>= λ a → 42 | nativeReturn a 43 | takeVar (addVar A v (addVar B v' l)) = nativeTakeVar {A} v native>>= λ a → 44 | takeVar (addVar B v' l) native>>= λ rest → 45 | nativeReturn ( a , rest ) 46 | 47 | 48 | 49 | putVar : (l : VarList) → prod l → NativeIO Unit 50 | putVar [] _ = nativeReturn _ 51 | putVar (addVar A v []) a = nativePutVar {A} v a 52 | putVar (addVar A v (addVar B v' l)) (a , rest) 53 | = nativePutVar {A} v a native>>= λ _ → 54 | putVar (addVar B v' l) rest native>>= 55 | nativeReturn 56 | 57 | dispatch : (l : VarList) → (prod l → NativeIO (prod l)) → NativeIO Unit 58 | dispatch l f = takeVar l native>>= λ a → 59 | f a native>>= λ a₁ → 60 | putVar l a₁ 61 | 62 | dispatchList : (l : VarList) → List (prod l → NativeIO (prod l)) → NativeIO Unit 63 | dispatchList l [] = nativeReturn _ 64 | dispatchList l (p ∷ rest) = dispatch l p native>>= λ _ → 65 | dispatchList l rest 66 | -------------------------------------------------------------------------------- /src/StateSizedIO/IOObject.agda: -------------------------------------------------------------------------------- 1 | module StateSizedIO.IOObject where 2 | 3 | open import Data.Product 4 | 5 | open import Size 6 | 7 | open import SizedIO.Base 8 | open import StateSizedIO.Object 9 | 10 | 11 | 12 | -- --- Moved to BaseNonPoly.agda 13 | -- --- and /src/StateSizedIO/GUI/BaseStateDependent.agda 14 | -- --- FILE IS DELETED !!! 15 | -- --- 16 | -- --- 17 | 18 | 19 | -- An IO object is like a simple object, 20 | -- but the method returns IO applied to the result type of a simple object 21 | -- which means the method returns an IO program which when terminating 22 | -- returns the result of the simple object 23 | 24 | 25 | 26 | {- NOTE IOObject is now replaced by IOObjectˢ as defined in 27 | StateSizedIO.Base 28 | -} 29 | 30 | {- 31 | module _ (ioi : IOInterface) (let C = Command ioi) (let R = Response ioi) 32 | (oi : Interfaceˢ) (let S = Stateˢ oi) (let M = Methodˢ oi) 33 | (let Rt = Resultˢ oi) (let next = nextˢ oi) 34 | 35 | where 36 | 37 | record IOObject (i : Size) (s : S) : Set where 38 | coinductive 39 | field 40 | method : ∀{j : Size< i} (m : M s) → IO ioi ∞ (Σ[ r ∈ Rt s m ] IOObject j (next s m r) ) 41 | 42 | open IOObject public 43 | -} 44 | -------------------------------------------------------------------------------- /src/StateSizedIO/Object.agda: -------------------------------------------------------------------------------- 1 | module StateSizedIO.Object where 2 | 3 | open import Data.Product 4 | open import Level renaming (zero to lzero; suc to lsuc) 5 | open import StateSizedIO.GUI.BaseStateDependent 6 | -- open import StateSizedIO.Base public 7 | {- 8 | -- This definition was probably moved to StateSizedIO.Base 9 | -- and by accident left here. Delete this. 10 | record Interfaceˢ : Set₁ where 11 | field 12 | Stateˢ : Set 13 | Methodˢ : Stateˢ → Set 14 | Resultˢ : (s : Stateˢ) → (m : Methodˢ s) → Set 15 | nextˢ : (s : Stateˢ) → (m : Methodˢ s) → Resultˢ s m → Stateˢ 16 | open Interfaceˢ public 17 | -} 18 | 19 | module _ (I : Interfaceˢ {lzero} {lzero} {lzero})(let S = Stateˢ I) (let M = Methodˢ I) 20 | (let R = Resultˢ I) (let next = nextˢ I) where 21 | -- A simple object just returns for a method the response 22 | -- and the object itself 23 | record Objectˢ (s : S) : Set where 24 | coinductive 25 | field 26 | objectMethod : (m : M s) → Σ[ r ∈ R s m ] Objectˢ (next s m r) 27 | 28 | open Objectˢ public 29 | 30 | 31 | 32 | _▹_ : {A : Set } → { B : Set } → A → (A → B) → B 33 | a ▹ f = f a 34 | 35 | 36 | -- Bisimilation for Objectˢ 37 | 38 | module Bisim (I : Interfaceˢ) 39 | (let S = Stateˢ I) 40 | (let M = Methodˢ I) 41 | (let R = Resultˢ I) 42 | (let next = nextˢ I) 43 | (let O = Objectˢ I) 44 | where 45 | 46 | mutual 47 | 48 | record _≅_ {s : S} (o o' : O s) : Set where 49 | coinductive 50 | field bisimMethod : (m : M s) → objectMethod o m ≡×≅ objectMethod o' m 51 | 52 | data _≡×≅_ {s m} : (p p' : Σ[ r ∈ R s m ] O (next s m r)) → Set where 53 | bisim : ∀{r} (let s' = next s m r) {o o' : O s'} (p : o ≅ o') 54 | → (r , o) ≡×≅ (r , o') 55 | 56 | open _≅_ public 57 | 58 | refl≅ : ∀{s} (o : O s) → o ≅ o 59 | bisimMethod (refl≅ o) m = bisim (refl≅ (proj₂ (objectMethod o m))) 60 | -------------------------------------------------------------------------------- /src/UnSizedIO/Base.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --postfix-projections #-} 2 | 3 | module UnSizedIO.Base where 4 | 5 | open import Data.Maybe.Base 6 | open import Data.Sum renaming (inj₁ to left; inj₂ to right; [_,_]′ to either) 7 | open import Function 8 | 9 | open import NativeIO 10 | 11 | record IOInterface : Set₁ where 12 | field 13 | Command : Set 14 | Response : (m : Command) → Set 15 | open IOInterface public 16 | 17 | module _ (I : IOInterface) (let C = Command I) (let R = Response I) (A : Set) 18 | where 19 | 20 | mutual 21 | record IO : Set where 22 | coinductive 23 | constructor delay 24 | field 25 | force : IO' 26 | 27 | data IO' : Set where 28 | exec' : (c : C) (f : R c → IO) → IO' 29 | return' : (a : A) → IO' 30 | 31 | open IO public 32 | 33 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) 34 | where 35 | return : ∀{A} (a : A) → IO I A 36 | return a .force = return' a 37 | 38 | exec : ∀{A} (c : C) (f : R c → IO I A) → IO I A 39 | exec c f .force = exec' c f 40 | 41 | exec1 : (c : C) → IO I (R c) 42 | exec1 c = exec c return 43 | 44 | infixl 2 _>>=_ _>>='_ _>>_ 45 | 46 | mutual 47 | _>>='_ : ∀{A B} (m : IO' I A) (k : A → IO I B) → IO' I B 48 | exec' c f >>=' k = exec' c λ x → f x >>= k 49 | return' a >>=' k = force (k a) 50 | 51 | _>>=_ : ∀{A B} (m : IO I A) (k : A → IO I B) → IO I B 52 | force (m >>= k) = force m >>=' k 53 | 54 | _>>_ : ∀{B} (m : IO I Unit) (k : IO I B) → IO I B 55 | m >> k = m >>= λ _ → k 56 | 57 | 58 | module _ {I : IOInterface} (let C = Command I) (let R = Response I) 59 | where 60 | {-# NON_TERMINATING #-} 61 | translateIO : ∀ {A} 62 | → (translateLocal : (c : C) → NativeIO (R c)) 63 | → IO I A 64 | → NativeIO A 65 | translateIO translateLocal m = case (m .force) of 66 | λ{ (exec' c f) → (translateLocal c) native>>= λ r → 67 | translateIO translateLocal (f r) 68 | ; (return' a) → nativeReturn a 69 | } 70 | 71 | -- Recursion 72 | 73 | -- trampoline provides a generic form of loop (generalizing while/repeat). 74 | -- Starting at state s : S, step function f is iterated until it returns 75 | -- a result in A. 76 | 77 | module _ (I : IOInterface)(let C = Command I) (let R = Response I) 78 | where 79 | data IO+ (A : Set) : Set where 80 | exec' : (c : C) (f : R c → IO I A) → IO+ A 81 | 82 | module _ {I : IOInterface}(let C = Command I) (let R = Response I) 83 | where 84 | fromIO+' : ∀{A} → IO+ I A → IO' I A 85 | fromIO+' (exec' c f) = exec' c f 86 | 87 | fromIO+ : ∀{A} → IO+ I A → IO I A 88 | fromIO+ (exec' c f) .force = exec' c f 89 | 90 | _>>=+'_ : ∀{A B} (m : IO+ I A) (k : A → IO I B) → IO' I B 91 | exec' c f >>=+' k = exec' c λ x → f x >>= k 92 | 93 | _>>=+_ : ∀{A B} (m : IO+ I A) (k : A → IO I B) → IO I B 94 | force (m >>=+ k) = m >>=+' k 95 | 96 | mutual 97 | 98 | _>>+_ : ∀{A B} (m : IO I (A ⊎ B)) (k : A → IO I B) → IO I B 99 | force (m >>+ k) = force m >>+' k 100 | 101 | _>>+'_ : ∀{A B} (m : IO' I (A ⊎ B)) (k : A → IO I B) → IO' I B 102 | exec' c f >>+' k = exec' c λ x → f x >>+ k 103 | return' (left a) >>+' k = force (k a) 104 | return' (right b) >>+' k = return' b 105 | 106 | -- loop 107 | 108 | {-# TERMINATING #-} 109 | trampoline : ∀{A S} (f : S → IO+ I (S ⊎ A)) (s : S) → IO I A 110 | force (trampoline f s) = case (f s) of 111 | \{ (exec' c k) → exec' c λ r → k r >>+ trampoline f } 112 | 113 | -- simple infinite loop 114 | 115 | {-# TERMINATING #-} 116 | forever : ∀{A B} → IO+ I A → IO I B 117 | force (forever (exec' c f)) = exec' c λ r → f r >>= λ _ → forever (exec' c f) 118 | 119 | whenJust : {A : Set} → Maybe A → (A → IO I Unit) → IO I Unit 120 | whenJust nothing k = return _ 121 | whenJust (just a) k = k a 122 | 123 | module _ (I : IOInterface ) 124 | (let C = I .Command) 125 | (let R = I .Response) 126 | where 127 | 128 | data IOind (A : Set) : Set where 129 | 130 | exec'' : (c : C) (f : (r : R c) → IOind A) → IOind A 131 | return'' : (a : A) → IOind A 132 | 133 | 134 | 135 | main : NativeIO Unit 136 | main = nativePutStrLn "Hello, world!" 137 | -------------------------------------------------------------------------------- /src/UnSizedIO/Console.agda: -------------------------------------------------------------------------------- 1 | module UnSizedIO.Console where 2 | 3 | open import NativeIO 4 | open import UnSizedIO.Base hiding (main) 5 | 6 | data ConsoleCommand : Set where 7 | putStrLn : String → ConsoleCommand 8 | getLine : ConsoleCommand 9 | 10 | ConsoleResponse : ConsoleCommand → Set 11 | ConsoleResponse (putStrLn s) = Unit 12 | ConsoleResponse getLine = String 13 | 14 | ConsoleInterface : IOInterface 15 | Command ConsoleInterface = ConsoleCommand 16 | Response ConsoleInterface = ConsoleResponse 17 | 18 | IOConsole : Set → Set 19 | IOConsole = IO ConsoleInterface 20 | 21 | IOConsole+ : Set → Set 22 | IOConsole+ = IO+ ConsoleInterface 23 | 24 | translateIOConsoleLocal : (c : ConsoleCommand) → NativeIO (ConsoleResponse c) 25 | translateIOConsoleLocal (putStrLn s) = nativePutStrLn s 26 | translateIOConsoleLocal getLine = nativeGetLine 27 | 28 | translateIOConsole : {A : Set} → IOConsole A → NativeIO A 29 | translateIOConsole = translateIO translateIOConsoleLocal 30 | 31 | main : NativeIO Unit 32 | main = nativePutStrLn "Console" 33 | -------------------------------------------------------------------------------- /src/UnSizedIO/ConsoleObject.agda: -------------------------------------------------------------------------------- 1 | module UnSizedIO.ConsoleObject where 2 | 3 | open import UnSizedIO.Console 4 | open import UnSizedIO.Object 5 | open import UnSizedIO.IOObject 6 | 7 | -- A console object is an IO object for the IO interface of console 8 | ConsoleObject : (iface : Interface) → Set 9 | ConsoleObject iface = IOObject ConsoleInterface iface 10 | 11 | 12 | 13 | -------------------------------------------------------------------------------- /src/UnSizedIO/IOObject.agda: -------------------------------------------------------------------------------- 1 | module UnSizedIO.IOObject where 2 | 3 | open import Data.Product 4 | 5 | open import UnSizedIO.Base 6 | open import UnSizedIO.Object 7 | 8 | -- An IO object is like a simple object, 9 | -- but the method returns IO applied to the result type of a simple object 10 | -- which means the method returns an IO program which when terminating 11 | -- returns the result of the simple object 12 | 13 | 14 | module _ (ioi : IOInterface) (let C = Command ioi) (let R = Response ioi) 15 | (oi : Interface) (let M = Method oi) (let Rt = Result oi) 16 | where 17 | 18 | record IOObject : Set where 19 | coinductive 20 | field 21 | method : (m : M) → IO ioi (Rt m × IOObject) 22 | 23 | open IOObject public 24 | -------------------------------------------------------------------------------- /src/UnSizedIO/Object.agda: -------------------------------------------------------------------------------- 1 | module UnSizedIO.Object where 2 | 3 | open import Data.Product 4 | 5 | 6 | record Interface : Set₁ where 7 | field 8 | Method : Set 9 | Result : (m : Method) → Set 10 | open Interface public 11 | 12 | -- A simple object just returns for a method the response 13 | -- and the object itself 14 | record Object (i : Interface) : Set where 15 | coinductive 16 | field 17 | objectMethod : (m : Method i) → Result i m × Object i 18 | 19 | open Object public 20 | -------------------------------------------------------------------------------- /src/Unit.agda: -------------------------------------------------------------------------------- 1 | 2 | module Unit where 3 | 4 | record Unit : Set where 5 | constructor unit 6 | 7 | {-# COMPILE GHC Unit = data () (()) #-} 8 | -------------------------------------------------------------------------------- /src/lib/libraryDec.agda: -------------------------------------------------------------------------------- 1 | module lib.libraryDec where 2 | 3 | -- old version of Dec 4 | -- 5 | 6 | open import Data.Empty 7 | open import Level 8 | open import Relation.Nullary using ( ¬_ ) 9 | 10 | data Dec {p} (P : Set p) : Set p where 11 | yes : ( p : P) → Dec P 12 | no : (¬p : ¬ P) → Dec P 13 | -------------------------------------------------------------------------------- /src/lib/libraryString.agda: -------------------------------------------------------------------------------- 1 | module lib.libraryString where 2 | 3 | open import Data.String 4 | open import Data.Bool 5 | open import Agda.Builtin.String using ( primStringEquality ) 6 | open import Data.String public using () 7 | renaming (toList to primStringToList; 8 | _++_ to _++Str_) 9 | 10 | _==Str_ : String → String → Bool 11 | _==Str_ = primStringEquality 12 | -------------------------------------------------------------------------------- /src/ooAgda.agda-lib: -------------------------------------------------------------------------------- 1 | name: ooAgda 2 | depend: standard-library 3 | include: . ../examples 4 | --------------------------------------------------------------------------------