├── .gitignore ├── .stylish-haskell.yaml ├── LICENSE ├── README.md ├── Setup.hs ├── agda-ocaml.cabal ├── app └── Main.hs ├── benchmark └── agda-ocaml │ ├── .gitignore │ ├── AgdaListGen.hs │ ├── Extra.agda │ ├── Fold0.agda │ ├── Main.hs │ ├── RbTyped │ ├── RbUntyped │ ├── RedBlack.agda │ ├── RedBlack.hs │ ├── RedBlackMlf │ ├── Typed.hs │ ├── TypedExist.hs │ ├── Untyped.hs │ ├── run.sh │ └── spinner.sh ├── src └── full │ └── Agda │ └── Compiler │ ├── Malfunction.hs │ └── Malfunction │ ├── AST.hs │ ├── Compiler.hs │ ├── Instances.hs │ ├── Main.hs │ ├── Primitive.hs │ └── Run.hs ├── stack-8.0.2.yaml └── test └── agda-ocaml ├── CompilerTest.hs ├── Golden ├── Constructor.agda ├── Constructor_a.golden ├── Constructor_one.golden ├── Factorial.agda ├── Factorial_a.golden ├── Factorial_b.golden ├── FstSnd.agda ├── FstSnd_a.golden ├── FstSnd_b.golden ├── Index.agda ├── Index_l0.golden ├── Index_l1.golden ├── Index_l2.golden ├── Insertion.agda ├── InsertionSort.agda ├── InsertionSort_l0.golden ├── InsertionSort_l1.golden ├── InsertionSort_l2.golden ├── Insertion_l0.golden ├── Insertion_l1.golden ├── RedBlack.agda └── RedBlack_short.golden ├── Tasty.hs └── Utils.hs /.gitignore: -------------------------------------------------------------------------------- 1 | # Keep this file in alphabetic order please! 2 | 3 | .*.sw[a-p] 4 | .*.vim 5 | .cabal-sandbox/ 6 | .stack-work/ 7 | .DS_Store 8 | .sw[a-p] 9 | *.agdai 10 | *.hi 11 | *.o 12 | *.tmp 13 | *.tmp.* 14 | *~ 15 | \#*\# 16 | .\#* 17 | agda-ffi/dist 18 | agda-stdlib 19 | autom4te.cache 20 | cabal.sandbox.config 21 | config.log 22 | config.status 23 | configure 24 | dist/ 25 | dist-*/ 26 | examples/compiler/main 27 | hlint-report.html 28 | test/**/MAlonzo 29 | examples/**/MAlonzo 30 | mk/config.mk 31 | mk/parallel-tests.mk 32 | notes/style/haskell-style.tex 33 | src/full/Agda/Syntax/Parser/Lexer.hs 34 | src/full/Agda/Syntax/Parser/Parser.hs 35 | src/full/tags 36 | src/full/TAGS 37 | stack.yaml 38 | test/compiler/Main 39 | test/hello 40 | test/latex-backend/succeed/latex/ 41 | test/succeed/latex/ 42 | test/succeed/Issue1441 43 | exec-test* 44 | pkg-build* 45 | -------------------------------------------------------------------------------- /.stylish-haskell.yaml: -------------------------------------------------------------------------------- 1 | # stylish-haskell configuration file 2 | # ================================== 3 | 4 | # The stylish-haskell tool is mainly configured by specifying steps. These steps 5 | # are a list, so they have an order, and one specific step may appear more than 6 | # once (if needed). Each file is processed by these steps in the given order. 7 | steps: 8 | # Convert some ASCII sequences to their Unicode equivalents. This is disabled 9 | # by default. 10 | # - unicode_syntax: 11 | # # In order to make this work, we also need to insert the UnicodeSyntax 12 | # # language pragma. If this flag is set to true, we insert it when it's 13 | # # not already present. You may want to disable it if you configure 14 | # # language extensions using some other method than pragmas. Default: 15 | # # true. 16 | # add_language_pragma: true 17 | 18 | # Align the right hand side of some elements. This is quite conservative 19 | # and only applies to statements where each element occupies a single 20 | # line. 21 | - simple_align: 22 | cases: true 23 | top_level_patterns: true 24 | records: true 25 | 26 | # Import cleanup 27 | - imports: 28 | # There are different ways we can align names and lists. 29 | # 30 | # - global: Align the import names and import list throughout the entire 31 | # file. 32 | # 33 | # - file: Like global, but don't add padding when there are no qualified 34 | # imports in the file. 35 | # 36 | # - group: Only align the imports per group (a group is formed by adjacent 37 | # import lines). 38 | # 39 | # - none: Do not perform any alignment. 40 | # 41 | # Default: global. 42 | align: global 43 | 44 | # Folowing options affect only import list alignment. 45 | # 46 | # List align has following options: 47 | # 48 | # - after_alias: Import list is aligned with end of import including 49 | # 'as' and 'hiding' keywords. 50 | # 51 | # > import qualified Data.List as List (concat, foldl, foldr, head, 52 | # > init, last, length) 53 | # 54 | # - with_alias: Import list is aligned with start of alias or hiding. 55 | # 56 | # > import qualified Data.List as List (concat, foldl, foldr, head, 57 | # > init, last, length) 58 | # 59 | # - new_line: Import list starts always on new line. 60 | # 61 | # > import qualified Data.List as List 62 | # > (concat, foldl, foldr, head, init, last, length) 63 | # 64 | # Default: after_alias 65 | list_align: after_alias 66 | 67 | # Long list align style takes effect when import is too long. This is 68 | # determined by 'columns' setting. 69 | # 70 | # - inline: This option will put as much specs on same line as possible. 71 | # 72 | # - new_line: Import list will start on new line. 73 | # 74 | # - new_line_multiline: Import list will start on new line when it's 75 | # short enough to fit to single line. Otherwise it'll be multiline. 76 | # 77 | # - multiline: One line per import list entry. 78 | # Type with contructor list acts like single import. 79 | # 80 | # > import qualified Data.Map as M 81 | # > ( empty 82 | # > , singleton 83 | # > , ... 84 | # > , delete 85 | # > ) 86 | # 87 | # Default: inline 88 | long_list_align: inline 89 | 90 | # List padding determines indentation of import list on lines after import. 91 | # This option affects 'list_align' and 'long_list_align'. 92 | list_padding: 4 93 | 94 | # Separate lists option affects formating of import list for type 95 | # or class. The only difference is single space between type and list 96 | # of constructors, selectors and class functions. 97 | # 98 | # - true: There is single space between Foldable type and list of it's 99 | # functions. 100 | # 101 | # > import Data.Foldable (Foldable (fold, foldl, foldMap)) 102 | # 103 | # - false: There is no space between Foldable type and list of it's 104 | # functions. 105 | # 106 | # > import Data.Foldable (Foldable(fold, foldl, foldMap)) 107 | # 108 | # Default: true 109 | separate_lists: true 110 | 111 | # Language pragmas 112 | - language_pragmas: 113 | # We can generate different styles of language pragma lists. 114 | # 115 | # - vertical: Vertical-spaced language pragmas, one per line. 116 | # 117 | # - compact: A more compact style. 118 | # 119 | # - compact_line: Similar to compact, but wrap each line with 120 | # `{-#LANGUAGE #-}'. 121 | # 122 | # Default: vertical. 123 | style: vertical 124 | 125 | # Align affects alignment of closing pragma brackets. 126 | # 127 | # - true: Brackets are aligned in same collumn. 128 | # 129 | # - false: Brackets are not aligned together. There is only one space 130 | # between actual import and closing bracket. 131 | # 132 | # Default: true 133 | align: true 134 | 135 | # stylish-haskell can detect redundancy of some language pragmas. If this 136 | # is set to true, it will remove those redundant pragmas. Default: true. 137 | remove_redundant: true 138 | 139 | # Replace tabs by spaces. This is disabled by default. 140 | # - tabs: 141 | # # Number of spaces to use for each tab. Default: 8, as specified by the 142 | # # Haskell report. 143 | # spaces: 8 144 | 145 | # Remove trailing whitespace 146 | - trailing_whitespace: {} 147 | 148 | # A common setting is the number of columns (parts of) code will be wrapped 149 | # to. Different steps take this into account. Default: 80. 150 | columns: 80 151 | 152 | # By default, line endings are converted according to the OS. You can override 153 | # preferred format here. 154 | # 155 | # - native: Native newline format. CRLF on Windows, LF on other OSes. 156 | # 157 | # - lf: Convert to LF ("\n"). 158 | # 159 | # - crlf: Convert to CRLF ("\r\n"). 160 | # 161 | # Default: native. 162 | newline: native 163 | 164 | # Sometimes, language extensions are specified in a cabal file or from the 165 | # command line instead of using language pragmas in the file. stylish-haskell 166 | # needs to be aware of these, so it can parse the file correctly. 167 | # 168 | language_extensions: 169 | - TypeOperators 170 | - FlexibleContexts 171 | - GADTs 172 | - RecordWildCards 173 | - ScopedTypeVariables 174 | - NamedFieldPuns 175 | - TupleSections 176 | - MultiWayIf 177 | - MultiParamTypeClasses 178 | - ScopedTypeVariables 179 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Copyright Jan Mas Rovira, Frederik Hanghøj Iversen (c) 2017 2 | 3 | All rights reserved. 4 | 5 | Redistribution and use in source and binary forms, with or without 6 | modification, are permitted provided that the following conditions are met: 7 | 8 | * Redistributions of source code must retain the above copyright 9 | notice, this list of conditions and the following disclaimer. 10 | 11 | * Redistributions in binary form must reproduce the above 12 | copyright notice, this list of conditions and the following 13 | disclaimer in the documentation and/or other materials provided 14 | with the distribution. 15 | 16 | * Neither the name of Jan Mas Rovira nor the names of other 17 | contributors may be used to endorse or promote products derived 18 | from this software without specific prior written permission. 19 | 20 | THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS 21 | "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT 22 | LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR 23 | A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT 24 | OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 25 | SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT 26 | LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, 27 | DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY 28 | THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT 29 | (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE 30 | OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. 31 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | Introduction 2 | ============ 3 | This is an experimental backend for Agda that targets the OCaml compiler. 4 | 5 | This repo is a mirror (and a branch off of) the repo that was used 6 | during development. The original repo can be found here: 7 | 8 | https://gitlab.com/janmasrovira/agda2mlf 9 | 10 | The major differece is the project structure. And the above repo also 11 | contains a report we wrote for the project. 12 | 13 | Project structure 14 | ================= 15 | The project layout is designed to mirror suitable locations in `agda`. 16 | 17 | Building 18 | ======== 19 | You need to install `malfunction` to get up and running. Please refer 20 | to [malfunction] to see how to get this on your system. 21 | 22 | This project has been tested using `stack`. Building should be as easy as: 23 | 24 | ln -s stack-8.0.2.yaml stack.yaml 25 | stack build 26 | 27 | Running: 28 | 29 | stack exec agda-ocaml 30 | 31 | Installing: 32 | 33 | stack install agda-ocaml 34 | 35 | Testing: 36 | 37 | stack test 38 | 39 | Benchmarking: 40 | 41 | stack exec agda-ocaml-bench 42 | 43 | The benchmarks depend on [agda-prelude] - so download that and add it to 44 | your default libraries. 45 | 46 | [malfunction]: https://github.com/stedolan/malfunction 47 | [agda-prelude]: https://github.com/UlfNorell/agda-prelude 48 | -------------------------------------------------------------------------------- /Setup.hs: -------------------------------------------------------------------------------- 1 | 2 | import Data.List 3 | import Distribution.Simple 4 | import Distribution.Simple.LocalBuildInfo 5 | import Distribution.Simple.Setup 6 | import Distribution.Simple.BuildPaths (exeExtension) 7 | import Distribution.PackageDescription 8 | import System.FilePath 9 | import System.Process 10 | import System.Exit 11 | 12 | main = defaultMainWithHooks hooks 13 | 14 | hooks = simpleUserHooks { regHook = checkAgdaPrimitiveAndRegister } 15 | 16 | builtins :: [String] 17 | builtins = 18 | [ "Bool", "Char", "Coinduction", "Equality", "Float" 19 | , "FromNat", "FromNeg", "FromString", "IO", "Int", "List" 20 | , "Nat", "Reflection", "Size", "Strict", "String" 21 | , "TrustMe", "Unit" ] 22 | 23 | checkAgdaPrimitive :: PackageDescription -> LocalBuildInfo -> RegisterFlags -> IO () 24 | checkAgdaPrimitive pkg info flags | regGenPkgConf flags /= NoFlag = return () -- Gets run twice, only do this the second time 25 | checkAgdaPrimitive pkg info flags = do 26 | let dirs = absoluteInstallDirs pkg info NoCopyDest 27 | agda = buildDir info "agda" "agda" <.> exeExtension 28 | primMod ms = (ms, datadir dirs "lib" "prim" "Agda" foldr1 () ms <.> "agda") 29 | prims = primMod ["Primitive"] : [ primMod ["Builtin", m] | m <- builtins ] 30 | 31 | checkPrim (ms, file) = do 32 | ok <- rawSystem agda [file, "-v0"] 33 | case ok of 34 | ExitSuccess -> return () 35 | ExitFailure _ -> putStrLn $ "WARNING: Failed to typecheck " ++ intercalate "." ("Agda" : ms) ++ "!" 36 | 37 | putStrLn "Generating Agda library interface files..." 38 | mapM_ checkPrim prims 39 | 40 | checkAgdaPrimitiveAndRegister :: PackageDescription -> LocalBuildInfo -> UserHooks -> RegisterFlags -> IO () 41 | checkAgdaPrimitiveAndRegister pkg info hooks flags = do 42 | checkAgdaPrimitive pkg info flags 43 | regHook simpleUserHooks pkg info hooks flags -- This actually does something useful 44 | 45 | -------------------------------------------------------------------------------- /agda-ocaml.cabal: -------------------------------------------------------------------------------- 1 | name: agda-ocaml 2 | version: 0.2.0.0 3 | synopsis: - 4 | description: Agda to Ocaml 5 | license: BSD3 6 | license-file: LICENSE 7 | author: Jan Mas Rovira, Frederik Hanghøj Iversen 8 | maintainer: janmasrovira@gmail.com, hanghj@student.chalmers.se 9 | copyright: Jan Mas Rovira, Frederik Hanghøj Iversen 10 | category: Compiler 11 | build-type: Simple 12 | extra-source-files: README.md 13 | cabal-version: >=1.10 14 | 15 | library 16 | hs-source-dirs: src/full 17 | exposed-modules: Agda.Compiler.Malfunction.AST 18 | , Agda.Compiler.Malfunction.Run 19 | , Agda.Compiler.Malfunction.Compiler 20 | , Agda.Compiler.Malfunction.Main 21 | , Agda.Compiler.Malfunction 22 | , Agda.Compiler.Malfunction.Primitive 23 | , Agda.Compiler.Malfunction.Instances 24 | build-depends: Agda >= 2.5.2 25 | , base >= 4.7 && < 5 26 | , containers == 0.5.* 27 | , extra == 1.5.* 28 | , filepath == 1.4.* 29 | , mtl == 2.2.* 30 | , pretty == 1.1.* 31 | , process == 1.4.* 32 | , temporary == 1.2.* 33 | , uniplate == 1.6.* 34 | default-language: Haskell2010 35 | 36 | executable agda-ocaml 37 | hs-source-dirs: app 38 | main-is: Main.hs 39 | ghc-options: -threaded -rtsopts -with-rtsopts=-N 40 | build-depends: base 41 | , agda-ocaml 42 | default-language: Haskell2010 43 | 44 | test-suite agda-ocaml-test 45 | type: exitcode-stdio-1.0 46 | hs-source-dirs: test/agda-ocaml 47 | other-modules: CompilerTest 48 | , Utils 49 | main-is: Tasty.hs 50 | build-depends: Agda 51 | , agda-ocaml 52 | , base 53 | , bytestring == 0.10.* 54 | , containers == 0.5.* 55 | , extra == 1.5.* 56 | , filepath == 1.4.* 57 | , tasty == 0.11.* 58 | , tasty-golden == 2.3.* 59 | , tasty-hunit == 0.9.* 60 | , directory == 1.3.* 61 | -- ghc-options: -threaded -rtsopts -with-rtsopts=-N 62 | default-language: Haskell2010 63 | 64 | executable agda-ocaml-bench 65 | hs-source-dirs: benchmark/agda-ocaml 66 | other-modules: RedBlack 67 | , Typed 68 | , TypedExist 69 | , Untyped 70 | , AgdaListGen 71 | main-is: Main.hs 72 | build-depends: 73 | base 74 | , directory == 1.3.* 75 | , process == 1.4.* 76 | default-language: Haskell2010 77 | 78 | 79 | source-repository head 80 | type: git 81 | location: git@github.com:agda/agda-ocaml.git 82 | -------------------------------------------------------------------------------- /app/Main.hs: -------------------------------------------------------------------------------- 1 | module Main (main) where 2 | 3 | import qualified Agda.Compiler.Malfunction.Main as Malfunction 4 | 5 | main :: IO () 6 | main = Malfunction.main 7 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/.gitignore: -------------------------------------------------------------------------------- 1 | RedBlack.mlf 2 | RedBlackMlf 3 | TheList.agda 4 | RedBlack 5 | RbUntyped 6 | RbTypedExist 7 | RbTyped 8 | Fold.agda 9 | RedBlackStrict 10 | HaskExistType 11 | HaskExistTypeStrict 12 | HaskTyped 13 | HaskTypedStrict 14 | HaskUntyped 15 | HaskUntypedStrict 16 | MAlonzoStrict 17 | MAlonzox 18 | Malfunction 19 | MalfunctionStrict 20 | RbTypedExistProf 21 | RbTypedPro 22 | RbUntypedProf 23 | RbUntypedT 24 | RbUntypedTail 25 | RedBlackAug 26 | RedBlackMlfAug 27 | RedBlackMlfStrict 28 | RedBlackT 29 | Fold.agda 30 | runs/ 31 | MAlonzo 32 | *.bench 33 | tmp.hs 34 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/AgdaListGen.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE CPP #-} 2 | module AgdaListGen where 3 | 4 | import Data.List 5 | 6 | theListLength :: Int 7 | theListLength = len 8 | 9 | #ifdef LSorted 10 | #ifdef tail 11 | theListDef = "theList = fromTo' " ++ show theListLength 12 | #else 13 | theListDef = "theList = fromTo " ++ show theListLength 14 | #endif 15 | #endif 16 | 17 | #ifdef LRSorted 18 | #ifdef tail 19 | theListDef = "theList = downFrom' " ++ show theListLength 20 | #else 21 | theListDef = "theList = downFrom " ++ show theListLength 22 | #endif 23 | #endif 24 | #ifdef LRandom 25 | #ifdef tail 26 | theListDef = "theList = downFromBbs' 43 " ++ show theListLength 27 | #else 28 | theListDef = "theList = fromToBbs 43 " ++ show theListLength 29 | #endif 30 | #endif 31 | 32 | main = putStrLn $ unlines ["open import Prelude" 33 | , "open import Extra" 34 | , "" 35 | , "theList : List Nat" 36 | , theListDef] 37 | 38 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/Extra.agda: -------------------------------------------------------------------------------- 1 | module Extra where 2 | 3 | open import Prelude 4 | 5 | 6 | fromTo : Nat -> List Nat 7 | fromTo n = f n 8 | where 9 | f : Nat -> List Nat 10 | f 0 = [] 11 | f (suc i) = n - i ∷ f i 12 | 13 | downFrom : Nat -> List Nat 14 | downFrom 0 = [] 15 | downFrom (suc n) = (suc n) ∷ downFrom n 16 | 17 | fromTo' : Nat -> List Nat 18 | fromTo' = f [] 19 | where 20 | f : List Nat -> Nat -> List Nat 21 | f xs 0 = xs 22 | f xs (suc x) = seq xs (f ( suc x ∷ xs ) x) 23 | 24 | downFrom' : Nat -> List Nat 25 | downFrom' n = f [] n 26 | where 27 | f : List Nat -> Nat -> List Nat 28 | f xs 0 = xs 29 | f xs (suc x) = seq xs (f ((n - x) ∷ xs) x) 30 | 31 | blumblumshub : Nat -> (m : Nat) -> {{ nz : NonZero m }} -> Nat 32 | blumblumshub xn M = natMod M (xn ^ 2) 33 | 34 | downFromBbs' : Nat -> Nat -> List Nat 35 | downFromBbs' seed = f seed [] 36 | where 37 | f : Nat -> List Nat -> Nat -> List Nat 38 | f _ acc 0 = acc 39 | f x acc (suc l) = seq acc $ f (blumblumshub x M) ( x ∷ acc ) l 40 | where 41 | M M17 M31 : Nat 42 | M = M17 * M31 43 | M17 = 2971 44 | M31 = 4111 45 | 46 | fromToBbs : Nat -> Nat -> List Nat 47 | fromToBbs _ 0 = [] 48 | fromToBbs xi (suc n) = xi ∷ fromToBbs (blumblumshub xi m) n 49 | where 50 | p = 2971 51 | q = 4111 52 | m = p * q 53 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/Fold0.agda: -------------------------------------------------------------------------------- 1 | module Fold where 2 | 3 | open import Prelude 4 | 5 | myfold : {ac b : Set} -> (ac -> b -> ac) -> ac -> List b -> ac 6 | #ifdef strict 7 | myfold = foldl! 8 | #else 9 | myfold = foldl 10 | #endif 11 | 12 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/Main.hs: -------------------------------------------------------------------------------- 1 | module Main where 2 | 3 | import System.Process 4 | import System.Directory 5 | 6 | main :: IO () 7 | main = withCurrentDirectory "benchmark/agda-ocaml" $ 8 | callProcess "./run.sh" [] 9 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/RbTyped: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/agda-ocaml/e38b699870ba638221828b07b12948d70a1bdaec/benchmark/agda-ocaml/RbTyped -------------------------------------------------------------------------------- /benchmark/agda-ocaml/RbUntyped: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/agda-ocaml/e38b699870ba638221828b07b12948d70a1bdaec/benchmark/agda-ocaml/RbUntyped -------------------------------------------------------------------------------- /benchmark/agda-ocaml/RedBlack.agda: -------------------------------------------------------------------------------- 1 | module RedBlack where 2 | 3 | open import Prelude hiding (insert) 4 | open import TheList 5 | open import Fold 6 | open import Extra 7 | 8 | -- Version of comparison that lets us use instance search for the proof objects. 9 | data Comparison! {a} {A : Set a} (_<_ : A → A → Set a) (x y : A) : Set a where 10 | less : {{_ : x < y}} → Comparison! _<_ x y 11 | equal : {{_ : x ≡ y}} → Comparison! _<_ x y 12 | greater : {{_ : y < x}} → Comparison! _<_ x y 13 | 14 | compare! : ∀ {a} {A : Set a} {{_ : Ord A}} (x y : A) → Comparison! _<_ x y 15 | compare! x y = 16 | case compare x y of λ where 17 | (less lt) → less {{lt}} 18 | (equal eq) → equal {{eq}} 19 | (greater gt) → greater {{gt}} 20 | 21 | record Prf (A : Set) : Set where 22 | constructor ! 23 | field 24 | {{prf}} : A 25 | 26 | data Bound (A : Set) : Set where 27 | -∞ ∞ : Bound A 28 | # : A → Bound A 29 | 30 | module _ {A : Set} {{_ : Ord A}} where 31 | LessBound : Bound A → Bound A → Set 32 | LessBound ∞ _ = ⊥ 33 | LessBound _ ∞ = ⊤ 34 | LessBound _ -∞ = ⊥ 35 | LessBound -∞ _ = ⊤ 36 | LessBound (# x) (# y) = x < y 37 | 38 | instance 39 | OrdBound : Ord (Bound A) 40 | OrdBound = defaultOrd compareBound 41 | where 42 | compareBound : (a b : Bound A) → Comparison LessBound a b 43 | compareBound -∞ -∞ = equal refl 44 | compareBound -∞ ∞ = less _ 45 | compareBound -∞ (# x) = less _ 46 | compareBound ∞ -∞ = greater tt 47 | compareBound ∞ ∞ = equal refl 48 | compareBound ∞ (# x) = greater tt 49 | compareBound (# x) -∞ = greater tt 50 | compareBound (# x) ∞ = less tt 51 | compareBound (# x) (# y) with compare x y 52 | ... | less lt = less lt 53 | ... | greater gt = greater gt 54 | compareBound (# x) (# .x) | equal refl = equal refl 55 | 56 | Bounds : Set → Set 57 | Bounds A = Bound A × Bound A 58 | 59 | Rel : Set → Set₁ 60 | Rel A = Bounds A → Set 61 | 62 | module _ {A : Set} where 63 | PrfR : Rel A → Rel A 64 | PrfR R b = Prf (R b) 65 | 66 | data _∧_ (S T : Rel A) : Rel A where 67 | pivot : ∀ {l u} p → S (l , # p) → T (# p , u) → (S ∧ T) (l , u) 68 | 69 | pattern -⟨_⟩- p = pivot p ! ! 70 | pattern _⟨_⟩_ l p r = pivot p l r 71 | 72 | module _ {A : Set} {{_ : Ord A}} where 73 | Less : Rel A 74 | Less = PrfR (uncurry _<_) 75 | 76 | Bounded : Rel A 77 | Bounded = Less ∧ Less 78 | 79 | data Color : Set where 80 | red black : Color 81 | 82 | module _ (A : Set) {{_ : Ord A}} where 83 | 84 | data Tree′ (b : Bounds A) : Nat → Color → Set 85 | 86 | Tree : Nat → Color → Rel A 87 | Tree n c b = Tree′ b n c 88 | 89 | data Tree′ b where 90 | leaf′ : Less b → Tree 0 black b 91 | red : ∀ {n} → (Tree n black ∧ Tree n black) b → Tree n red b 92 | black : ∀ {c₁ c₂ n} → (Tree n c₁ ∧ Tree n c₂) b → Tree (suc n) black b 93 | 94 | pattern leaf = leaf′ ! 95 | 96 | data PreTree′ (b : Bounds A) (n : Nat) : Color → Set 97 | 98 | PreTree : Nat → Color → Bounds A → Set 99 | PreTree n c b = PreTree′ b n c 100 | 101 | data PreTree′ b n where 102 | red : ∀ {c₁ c₂} → (Tree n c₁ ∧ Tree n c₂) b → PreTree n red b 103 | maybe-black : ∀ {c} → Tree n c b → PreTree n black b 104 | 105 | pattern _b⟨_⟩_ l x r = black (l ⟨ x ⟩ r) 106 | pattern _r⟨_⟩_ l x r = red (l ⟨ x ⟩ r) 107 | pattern _pb⟨_⟩_ l x r = maybe-black (black (l ⟨ x ⟩ r)) 108 | pattern _pbr⟨_⟩_ l x r = maybe-black (red (l ⟨ x ⟩ r)) 109 | pattern _rbb⟨_⟩_ l x r = red {c₁ = black} {c₂ = black} (l ⟨ x ⟩ r) 110 | 111 | module _ {A : Set} {{_ : Ord A}} where 112 | 113 | balance-l : ∀ {b c₁ c₂ n} → 114 | (PreTree A n c₁ ∧ Tree A n c₂) b → 115 | PreTree A (suc n) black b 116 | balance-l (((l₁ r⟨ z ⟩ l₂) r⟨ x ⟩ l₃) ⟨ y ⟩ r) = 117 | (l₁ b⟨ z ⟩ l₂) pbr⟨ x ⟩ (l₃ b⟨ y ⟩ r) 118 | balance-l ((l₁ r⟨ z ⟩ (l₂ r⟨ x ⟩ l₃)) ⟨ y ⟩ r) = 119 | (l₁ b⟨ z ⟩ l₂) pbr⟨ x ⟩ (l₃ b⟨ y ⟩ r) 120 | balance-l ((l₁ rbb⟨ x ⟩ l₂) ⟨ y ⟩ r) = 121 | (l₁ r⟨ x ⟩ l₂) pb⟨ y ⟩ r 122 | balance-l (maybe-black l ⟨ y ⟩ r) = l pb⟨ y ⟩ r 123 | 124 | balance-r : ∀ {b c₁ c₂ n} → 125 | (Tree A n c₁ ∧ PreTree A n c₂) b → 126 | PreTree A (suc n) black b 127 | balance-r (l ⟨ y ⟩ ((r₁ r⟨ z ⟩ r₂) r⟨ x ⟩ r₃)) = 128 | (l b⟨ y ⟩ r₁) pbr⟨ z ⟩ (r₂ b⟨ x ⟩ r₃) 129 | balance-r (l ⟨ y ⟩ (r₁ r⟨ x ⟩ (r₂ r⟨ z ⟩ r₃))) = 130 | (l b⟨ y ⟩ r₁) pbr⟨ x ⟩ (r₂ b⟨ z ⟩ r₃) 131 | balance-r (l ⟨ y ⟩ (r₁ rbb⟨ x ⟩ r₂)) = 132 | l pb⟨ y ⟩ (r₁ r⟨ x ⟩ r₂) 133 | balance-r (l ⟨ y ⟩ maybe-black r) = l pb⟨ y ⟩ r 134 | 135 | ins : ∀ {b c n} → Bounded b → Tree A n c b → PreTree A n c b 136 | ins -⟨ x ⟩- leaf = leaf pbr⟨ x ⟩ leaf 137 | ins -⟨ x ⟩- (red (l ⟨ y ⟩ r)) = 138 | case compare! x y of λ where 139 | less → case ins -⟨ x ⟩- l of λ { (maybe-black l′) → l′ r⟨ y ⟩ r } 140 | greater → case ins -⟨ x ⟩- r of λ { (maybe-black r′) → l r⟨ y ⟩ r′ } 141 | equal → l r⟨ y ⟩ r 142 | ins -⟨ x ⟩- (black (l ⟨ y ⟩ r)) = 143 | case compare! x y of λ where 144 | less → balance-l (ins -⟨ x ⟩- l ⟨ y ⟩ r) 145 | greater → balance-r (l ⟨ y ⟩ ins -⟨ x ⟩- r) 146 | equal → l pb⟨ y ⟩ r 147 | 148 | data RedBlackTree (A : Set) {{_ : Ord A}} : Set where 149 | mkT : ∀ {n} → Tree A n black (-∞ , ∞) → RedBlackTree A 150 | 151 | module _ {A : Set} {{_ : Ord A}} where 152 | 153 | insert : A → RedBlackTree A → RedBlackTree A 154 | insert x (mkT t) with ins -⟨ x ⟩- t 155 | ... | l pbr⟨ y ⟩ r = mkT $ l b⟨ y ⟩ r 156 | ... | l pb⟨ y ⟩ r = mkT $ l b⟨ y ⟩ r 157 | ... | maybe-black leaf = mkT leaf 158 | 159 | fromList : List A → RedBlackTree A 160 | -- we changed `foldr` here to `foldl!` for efficiency reasons 161 | fromList = myfold (flip insert) (mkT leaf) 162 | 163 | toList′ : ∀ {b n c} → Tree′ A b n c → List A → List A 164 | toList′ (leaf′ _) xs = xs 165 | toList′ (l r⟨ x ⟩ r) xs = toList′ l (x ∷ toList′ r xs) 166 | toList′ (l b⟨ x ⟩ r) xs = toList′ l (x ∷ toList′ r xs) 167 | 168 | toList : RedBlackTree A → List A 169 | toList (mkT t) = toList′ t [] 170 | 171 | treeSort : List A → List A 172 | treeSort = toList ∘ fromList 173 | 174 | ------------- 175 | -- Example -- 176 | ------------- 177 | 178 | main : IO Unit 179 | main = putStrLn $ show $ sum $ theList 180 | -- main = putStrLn $ show $ sum $ treeSort theList 181 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/RedBlack.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE CPP #-} 2 | {-# LANGUAGE BangPatterns #-} 3 | {-# LANGUAGE FlexibleInstances #-} 4 | 5 | module RedBlack where 6 | 7 | import Data.List (foldl') 8 | import Extra 9 | 10 | #ifdef RbUntyped 11 | import Untyped 12 | #endif 13 | #ifdef RbTyped 14 | import Typed 15 | #endif 16 | #ifdef RbTypedExist 17 | import TypedExist 18 | #endif 19 | 20 | len :: Int 21 | len = Len 22 | 23 | theList :: [Int] 24 | #ifdef LSorted 25 | #ifdef tail 26 | theList = fromTo' len 27 | #else 28 | theList = fromTo len 29 | #endif 30 | #endif 31 | #ifdef LRSorted 32 | #ifdef tail 33 | theList = downFrom' len 34 | #else 35 | theList = downFrom len 36 | #endif 37 | #endif 38 | #ifdef LRandom 39 | #ifdef tail 40 | theList = downFromBbs' 43 len 41 | #else 42 | theList = fromToBbs 43 len 43 | #endif 44 | #endif 45 | 46 | myFromList :: Ord a => [a] -> Tree a 47 | #ifdef strict 48 | myFromList = fromList' 49 | #else 50 | myFromList = fromList 51 | #endif 52 | 53 | 54 | -- We don't actually need to calculate the sum for the benchmark if we just 55 | -- redirect the output to `/dev/null` the IO will not be an overhead. It might 56 | -- actually be better *not* to calculate the sum because the `Foldable` 57 | -- instances (that I defined) may be more or less efficient. 58 | treeSort :: Ord a => [a] -> [a] 59 | treeSort = toList . myFromList 60 | 61 | bench :: Int 62 | bench = sum' $ treeSort theList 63 | -- bench = sum' $ theList 64 | 65 | sum' :: [Int] -> Int 66 | sum' = foldl' (+) 0 67 | 68 | main :: IO () 69 | main = print bench 70 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/RedBlackMlf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/agda/agda-ocaml/e38b699870ba638221828b07b12948d70a1bdaec/benchmark/agda-ocaml/RedBlackMlf -------------------------------------------------------------------------------- /benchmark/agda-ocaml/Typed.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE UndecidableInstances #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | module Typed (Tree, fromList, fromList', toList) where 4 | 5 | import qualified Data.Foldable as F 6 | import Data.List (foldl') 7 | 8 | {- Version 2, 1st typed version -} 9 | data Unit a = E deriving Show 10 | type Tr t a = (t a,a,t a) 11 | data Red t a = C (t a) | R (Tr t a) 12 | 13 | {- explicit Show instance as we work with 3rd order type constructors -} 14 | instance (Show (t a), Show a) => Show (Red t a) 15 | where showsPrec _ (C t) = shows t 16 | showsPrec _ (R(a,b,c)) = 17 | ("R("++) . shows a . (","++) . shows b . (","++) . shows c . (")"++) 18 | 19 | data AddLayer t a = B(Tr(Red t) a) deriving Show 20 | 21 | data RB t a = Base (t a) | Next (RB (AddLayer t) a) 22 | 23 | {- this Show instance is not Haskell98, but hugs -98 accepts it -} 24 | instance (Show (t a),Show a) => Show (RB t a) 25 | where 26 | show (Base t) = show t 27 | show (Next t) = show t 28 | 29 | type Tree a = RB Unit a 30 | empty :: Tree a 31 | empty = Base E 32 | 33 | type RR t a = Red (Red t) a 34 | type RL t a = Red (AddLayer t) a 35 | 36 | insert :: Ord a => a -> Tree a -> Tree a 37 | insert = rbinsert 38 | 39 | class Insertion t where ins :: Ord a => a -> t a -> Red t a 40 | instance Insertion Unit where ins x E = R(E,x,E) 41 | 42 | rbinsert :: (Ord a,Insertion t) => a -> RB t a -> RB t a 43 | rbinsert x (Next t) = Next (rbinsert x t) 44 | rbinsert x (Base t) = blacken(ins x t) 45 | 46 | blacken :: Red t a -> RB t a 47 | blacken (C u) = Base u 48 | blacken (R(a,x,b)) = Next(Base(B(C a,x,C b))) 49 | 50 | balanceL :: RR t a -> a -> Red t a -> RL t a 51 | balanceL (R(R(a,x,b),y,c)) z d = R(B(C a,x,C b),y,B(c,z,d)) 52 | balanceL (R(a,x,R(b,y,c))) z d = R(B(a,x,C b),y,B(C c,z,d)) 53 | balanceL (R(C a,x,C b)) z d = C(B(R(a,x,b),z,d)) 54 | balanceL (C a) x b = C(B(a,x,b)) 55 | 56 | balanceR :: Red t a -> a -> RR t a -> RL t a 57 | balanceR a x (R(R(b,y,c),z,d)) = R(B(a,x,C b),y,B(C c,z,d)) 58 | balanceR a x (R(b,y,R(c,z,d))) = R(B(a,x,b),y,B(C c,z,C d)) 59 | balanceR a x (R(C b,y,C c)) = C(B(a,x,R(b,y,c))) 60 | balanceR a x (C b) = C(B(a,x,b)) 61 | 62 | instance Insertion t => Insertion (AddLayer t) where 63 | ins x t@(B(l,y,r)) 64 | | xy = balance(C l) y (ins x r) 66 | | otherwise = C t 67 | instance Insertion t => Insertion (Red t) where 68 | ins x (C t) = C(ins x t) 69 | ins x t@(R(a,y,b)) 70 | | xy = R(C a,y,ins x b) 72 | | otherwise = C t 73 | 74 | balance :: RR t a -> a -> RR t a -> RL t a 75 | balance (R a) y (R b) = R(B a,y,B b) 76 | balance (C a) x b = balanceR a x b 77 | balance a x (C b) = balanceL a x b 78 | 79 | class Append t where app :: t a -> t a -> Red t a 80 | 81 | instance Append Unit where app _ _ = C E 82 | 83 | instance Append t => Append (AddLayer t) where 84 | app (B(a,x,b)) (B(c,y,d)) = threeformB a x (appRed b c) y d 85 | 86 | threeformB :: Red t a -> a -> RR t a -> a -> Red t a -> RL t a 87 | threeformB a x (R(b,y,c)) z d = R(B(a,x,b),y,B(c,z,d)) 88 | threeformB a x (C b) y c = balleftB (C a) x (B(b,y,c)) 89 | 90 | appRed :: Append t => Red t a -> Red t a -> RR t a 91 | appRed (C x) (C y) = C(app x y) 92 | appRed (C t) (R(a,x,b)) = R(app t a,x,C b) 93 | appRed (R(a,x,b)) (C t) = R(C a,x,app b t) 94 | appRed (R(a,x,b))(R(c,y,d)) = threeformR a x (app b c) y d 95 | 96 | threeformR:: t a -> a -> Red t a -> a -> t a -> RR t a 97 | threeformR a x (R(b,y,c)) z d = R(R(a,x,b),y,R(c,z,d)) 98 | threeformR a x (C b) y c = R(R(a,x,b),y,C c) 99 | 100 | balleft :: RR t a -> a -> RL t a -> RR (AddLayer t) a 101 | balleft (R a) y c = R(C(B a),y,c) 102 | balleft (C t) x (R(B(a,y,b),z,c)) = R(C(B(t,x,a)),y,balleftB (C b) z c) 103 | balleft b x (C t) = C (balleftB b x t) 104 | 105 | balleftB :: RR t a -> a -> AddLayer t a -> RL t a 106 | balleftB bl x (B y) = balance bl x (R y) 107 | 108 | balright :: RL t a -> a -> RR t a -> RR (AddLayer t) a 109 | balright a x (R b) = R(a,x,C(B b)) 110 | balright (R(a,x,B(b,y,c))) z (C d) = R(balrightB a x (C b),y,C(B(c,z,d))) 111 | balright (C t) x b = C (balrightB t x b) 112 | 113 | balrightB :: AddLayer t a -> a -> RR t a -> RL t a 114 | balrightB (B y) x t = balance (R y) x t 115 | 116 | class Append t => DelRed t where 117 | delTup :: Ord a => a -> Tr t a -> Red t a 118 | delLeft :: Ord a => a -> t a -> a -> Red t a -> RR t a 119 | delRight :: Ord a => a -> Red t a -> a -> t a -> RR t a 120 | 121 | class Append t => Del t where 122 | del :: Ord a => a -> AddLayer t a -> RR t a 123 | 124 | class (DelRed t, Del t) => Deletion t 125 | 126 | instance DelRed Unit where 127 | delTup z t@(_,x,_) = if x==z then C E else R t 128 | delLeft x _ y b = R(C E,y,b) 129 | delRight x a y _ = R(a,y,C E) 130 | 131 | instance Deletion t => DelRed (AddLayer t) where 132 | delTup z (a,x,b) 133 | | zx = balrightB a x (del z b) 135 | | otherwise = app a b 136 | delLeft x a y b = balleft (del x a) y b 137 | delRight x a y b = balright a y (del x b) 138 | 139 | instance (Append t, DelRed t) => Del t where 140 | del z (B(a,x,b)) 141 | | zx = delformRight b 143 | | otherwise = appRed a b 144 | where delformLeft(C t) = delLeft z t x b 145 | delformLeft(R t) = R(delTup z t,x,b) 146 | delformRight(C t) = delRight z a x t 147 | delformRight(R t) = R(a,x,delTup z t) 148 | 149 | instance Deletion t => Deletion (AddLayer t) 150 | instance Deletion Unit 151 | 152 | fromList' :: Ord a => [a] -> Tree a 153 | fromList' = foldl' (flip insert) empty 154 | 155 | fromList :: Ord a => [a] -> Tree a 156 | fromList = foldr insert empty 157 | 158 | 159 | instance Foldable Unit where 160 | foldr _ x E = x 161 | 162 | instance Foldable (RB Unit) where 163 | foldr f x t = case t of 164 | (Base E) -> x 165 | (Next t') -> foldr f x t' 166 | 167 | instance (Foldable a) => Foldable (RB (AddLayer a)) where 168 | foldr f x t = case t of 169 | (Base (B (lhs, a, rhs))) -> r 170 | where 171 | l = foldr f (f a x) lhs 172 | r = foldr f l rhs 173 | (Next t') -> foldr f x t' 174 | 175 | instance (Foldable a) => Foldable (AddLayer a) where 176 | foldr f x (B (lhs, a, rhs)) = r 177 | where 178 | l = foldr f (f a x) lhs 179 | r = foldr f l rhs 180 | 181 | instance Foldable a => Foldable (Red a) where 182 | foldr f x t = case t of 183 | C t' -> foldr f x t' 184 | R (lhs, a, rhs) -> r 185 | where 186 | l = foldr f (f a x) lhs 187 | r = foldr f l rhs 188 | 189 | toList :: Tree a -> [a] 190 | toList = F.toList 191 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/TypedExist.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE DeriveFoldable #-} 2 | {-# LANGUAGE StandaloneDeriving #-} 3 | {-# LANGUAGE ExistentialQuantification #-} 4 | module TypedExist (Tree, fromList, fromList', toList) where 5 | {- 3rd Version, 2nd typed version -} 6 | {- note that the type variable b is never used anywhere, only passed on. -} 7 | 8 | import Data.List (foldl') 9 | 10 | type Tr t a b = (t a b, a, t a b) 11 | data Red t a b = C (t a b) | R (Tr t a b) 12 | data Black a b = E | B (Tr (Red Black) a [b]) 13 | 14 | instance Show a => Show (Black a b) where 15 | showsPrec _ E = ('E':) 16 | showsPrec _ (B(a,x,b)) = ("B("++) . showRed a . (","++) . shows x . 17 | (","++) . showRed b . (")"++) 18 | 19 | showRed :: (Show (t1 a t), Show a) => Red t1 a t -> ShowS 20 | showRed (C x) = shows x 21 | showRed (R(a,x,b)) = ("R("++) . shows a . (","++) . shows x . (","++) . 22 | shows b . (")"++) 23 | 24 | type RR a b = Red (Red Black) a b 25 | 26 | inc :: Black a b -> Black a [b] 27 | inc = tickB 28 | 29 | {- tickB is the identity, 30 | but it allows us to replace that bogus type variable -} 31 | 32 | tickB :: Black a b -> Black a c 33 | tickB E = E 34 | tickB (B(a,x,b)) = B(tickR a,x,tickR b) 35 | 36 | tickR :: Red Black a b -> Red Black a c 37 | tickR (C t) = C(tickB t) 38 | tickR (R(a,x,b)) = R(tickB a,x,tickB b) 39 | 40 | data Tree a = forall b . ENC (Black a b) 41 | 42 | instance Show a => Show (Tree a) 43 | where show (ENC t) = show t 44 | 45 | empty :: Tree a 46 | empty = ENC E 47 | 48 | insert :: Ord a => a -> Tree a -> Tree a 49 | insert x (ENC t) = ENC(blacken (insB x t)) 50 | 51 | blacken :: Red Black a b -> Black a b 52 | blacken (C u) = u 53 | blacken (R(a,x,b)) = B(C(inc a),x,C(inc b)) 54 | 55 | insB :: Ord a => a -> Black a b -> Red Black a b 56 | insB x E = R(E,x,E) 57 | insB x t@(B(a,y,b)) 58 | | xy = balanceR a y (insR x b) 60 | | otherwise = C t 61 | 62 | insR :: Ord a => a -> Red Black a b -> RR a b 63 | insR x (C t) = C(insB x t) 64 | insR x t@(R(a,y,b)) 65 | | xy = R(C a,y,insB x b) 67 | | otherwise = C t 68 | 69 | --balanceL :: RR a [b] -> a -> Red Black a [b] -> Red Black a b 70 | balanceL (R(R(a,x,b),y,c)) z d = R(B(C a,x,C b),y,B(c,z,d)) 71 | balanceL (R(a,x,R(b,y,c))) z d = R(B(a,x,C b),y,B(C c,z,d)) 72 | balanceL (R(C a,x,C b)) z d = C(B(R(a,x,b),z,d)) 73 | balanceL (C a) x b = C(B(a,x,b)) 74 | 75 | --balanceR :: Red Black a [b] -> a -> RR a [b] -> Red Black a b 76 | balanceR a x (R(R(b,y,c),z,d)) = R(B(a,x,C b),y,B(C c,z,d)) 77 | balanceR a x (R(b,y,R(c,z,d))) = R(B(a,x,b),y,B(C c,z,C d)) 78 | balanceR a x (R(C b,y,C c)) = C(B(a,x,R(b,y,c))) 79 | balanceR a x (C b) = C(B(a,x,b)) 80 | 81 | 82 | fromList' :: Ord a => [a] -> Tree a 83 | fromList' = foldl' (flip insert) empty 84 | 85 | fromList :: Ord a => [a] -> Tree a 86 | fromList = foldr insert empty 87 | 88 | toList :: Tree a -> [a] 89 | toList (ENC tr) = toListBlack tr 90 | 91 | toListRed :: Red Black a [b] -> [a] 92 | toListRed (C rd) = toListBlack rd 93 | toListRed (R (l, a, r)) = toListBlack l ++ [a] ++ toListBlack r 94 | 95 | toListBlack :: Black a b -> [a] 96 | toListBlack E = [] 97 | toListBlack (B (l, a, r)) = toListRed l ++ [a] ++ toListRed r 98 | 99 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/Untyped.hs: -------------------------------------------------------------------------------- 1 | module Untyped (Tree, fromList, fromList', toList) where 2 | 3 | import qualified Data.Foldable as F 4 | import Data.List (foldl') 5 | 6 | {- Version 1, 'untyped' -} 7 | data Color = R | B deriving Show 8 | data RB a = E | T { color :: Color , left :: RB a , rootLabel :: a , right :: RB a } deriving Show 9 | 10 | {- Insertion and membership test as by Okasaki -} 11 | insert :: Ord a => a -> RB a -> RB a 12 | insert x s = 13 | T B a z b 14 | where 15 | T _ a z b = ins s 16 | ins E = T R E x E 17 | ins s@(T B a y b) 18 | | xy = balance a y (ins b) 20 | | otherwise = s 21 | ins s@(T R a y b) 22 | | xy = T R a y (ins b) 24 | | otherwise = s 25 | 26 | {- balance: first equation is new, 27 | to make it work with a weaker invariant -} 28 | balance :: RB a -> a -> RB a -> RB a 29 | balance (T R a x b) y (T R c z d) = T R (T B a x b) y (T B c z d) 30 | balance (T R (T R a x b) y c) z d = T R (T B a x b) y (T B c z d) 31 | balance (T R a x (T R b y c)) z d = T R (T B a x b) y (T B c z d) 32 | balance a x (T R b y (T R c z d)) = T R (T B a x b) y (T B c z d) 33 | balance a x (T R (T R b y c) z d) = T R (T B a x b) y (T B c z d) 34 | balance a x b = T B a x b 35 | 36 | type Tree a = RB a 37 | 38 | empty :: Tree a 39 | empty = E 40 | 41 | fromList' :: Ord a => [a] -> Tree a 42 | fromList' = foldl' (flip insert) empty 43 | 44 | fromList :: Ord a => [a] -> Tree a 45 | fromList = foldr insert empty 46 | 47 | instance Foldable RB where 48 | foldr f x t = case t of 49 | E{} -> x 50 | T{} -> r 51 | where 52 | a = rootLabel t 53 | l = foldr f (f a x) (left t) 54 | r = foldr f l (right t) 55 | 56 | toList :: Tree a -> [a] 57 | toList = F.toList 58 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/run.sh: -------------------------------------------------------------------------------- 1 | #!/bin/bash 2 | 3 | # this script is ugly and could use a lot of cleanup 4 | 5 | # lengths=`seq 250000 750000 30000000` 6 | lengths=`seq 1000000 2500000 50000000` 7 | execs="MAlonzox Malfunction" 8 | order_types="LRandom" 9 | 10 | date=`date` 11 | datedash=${date// /-} 12 | runsdir="runs" 13 | outdir="$runsdir/$datedash" 14 | mkdir -p $outdir 15 | echo $outdir 16 | lengths_file=$outdir/lengths.txt 17 | for len in $lengths 18 | do 19 | echo $len >> $lengths_file 20 | done 21 | 22 | time_cmd='/usr/bin/time --format "%e %M"' 23 | 24 | # $1 is the name of the executable 25 | # $2 is the order of the list 26 | getOutFile () { 27 | echo "$outdir/$1-$2.bench" 28 | } 29 | 30 | # $1 is the name of the executable 31 | # $2 is the order of the list 32 | bench () { 33 | outfile=$(getOutFile $1 $2) 34 | eval "$time_cmd ./$1 > /dev/null 2>> $outfile" 35 | } 36 | 37 | source ./spinner.sh 38 | for order in $order_types 39 | do 40 | echo "--------------------" 41 | echo "order = $order" 42 | for len in $lengths 43 | do 44 | echo "length = $len" 45 | 46 | 47 | start_spinner "Compiling MAlonzo" 48 | cpp AgdaListGen.hs -Dtail -Dlen=$len -D$order > tmp.hs 49 | stack exec runhaskell -- tmp.hs > TheList.agda 50 | cpp Fold0.agda Fold.agda 51 | sed '/^#/ d' Fold.agda -i 52 | stack exec agda-ocaml -- -c RedBlack.agda > /dev/null 53 | cp RedBlack MAlonzox 54 | stop_spinner $? 55 | 56 | # start_spinner "Compiling MAlonzo Strict" 57 | # cpp AgdaListGen.hs -Dtail -Dlen=$len -D$order > tmp.hs 58 | # stack exec runhaskell -- tmp.hs > TheList.agda 59 | # cpp Fold0.agda -Dstrict Fold.agda 60 | # sed '/^#/ d' Fold.agda -i 61 | # stack exec agda-ocaml -- -c RedBlack.agda > /dev/null 62 | # cp RedBlack MAlonzoStrict 63 | # stop_spinner $? 64 | 65 | # start_spinner "Compiling Mlf Strict" 66 | # cpp AgdaListGen.hs -Dtail -Dlen=$len -D$order > tmp.hs 67 | # stack exec runhaskell -- tmp.hs > TheList.agda 68 | # cpp Fold0.agda -Dtail -Dstrict Fold.agda 69 | # sed '/^#/ d' Fold.agda -i 70 | # stack exec agda-ocaml -- --mlf RedBlack.agda --compilemlf=MalfunctionStrict -o RedBlack.mlf > /dev/null 71 | # stop_spinner $? 72 | 73 | start_spinner "Compiling Mlf" 74 | cpp AgdaListGen.hs -Dtail -Dlen=$len -D$order > tmp.hs 75 | stack exec runhaskell -- tmp.hs > TheList.agda 76 | cpp Fold0.agda -Dtail Fold.agda 77 | sed '/^#/ d' Fold.agda -i 78 | stack exec agda-ocaml -- --mlf RedBlack.agda --compilemlf=Malfunction -o RedBlack.mlf > /dev/null 79 | stop_spinner $? 80 | 81 | 82 | # start_spinner "Compiling Haskell programs" 83 | # # ghc RedBlack.hs -O2 -prof -fprof-auto -rtsopts -main-is RedBlack -DRbTypedExist -D$order -DLen=$len -o RbTypedExistProf > /dev/null 84 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbUntyped -D$order -DLen=$len -o HaskUntyped > /dev/null 85 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbTyped -D$order -DLen=$len -o HaskTyped > /dev/null 86 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbTypedExist -D$order -DLen=$len -o HaskExistType > /dev/null 87 | # stop_spinner $? 88 | 89 | # start_spinner "Compiling Strict Haskell programs" 90 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbUntyped -Dstrict -D$order -DLen=$len -o HaskUntypedStrict > /dev/null 91 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbTyped -Dstrict -D$order -DLen=$len -o HaskTypedStrict > /dev/null 92 | # ghc RedBlack.hs -O2 -main-is RedBlack -DRbTypedExist -Dstrict -D$order -DLen=$len -o HaskExistTypeStrict > /dev/null 93 | # stop_spinner $? 94 | 95 | for exec in $execs 96 | do 97 | start_spinner "Running $exec, length = $len, order = $order" 98 | bench $exec $order 99 | stop_spinner $? 100 | done 101 | done 102 | done 103 | 104 | headfile=$outdir/head.txt 105 | printf "lengths" > $headfile 106 | for exec in $execs 107 | do 108 | for meas in time memory 109 | do 110 | printf " $exec-$meas" >> $headfile 111 | done 112 | done 113 | printf "\n" >> $headfile 114 | tablefile=$outdir/table.tmp 115 | for order in $order_types 116 | do 117 | cat $lengths_file > $tablefile 118 | for exec in $execs 119 | do 120 | file=$(getOutFile $exec $order) 121 | paste -d " " $tablefile $file > tmp 122 | mv tmp $tablefile 123 | done 124 | allfile=$outdir/$order.data 125 | cp $headfile $allfile 126 | cat $tablefile >> $allfile 127 | printf "\nBenchmark for $order\n" 128 | cat $allfile 129 | cp plot.plt $outdir 130 | (cd $outdir && gnuplot plot.plt) 131 | echo "" 132 | done 133 | -------------------------------------------------------------------------------- /benchmark/agda-ocaml/spinner.sh: -------------------------------------------------------------------------------- 1 | #!/bin/bash 2 | 3 | # Author: Tasos Latsas 4 | 5 | # spinner.sh 6 | # 7 | # Display an awesome 'spinner' while running your long shell commands 8 | # 9 | # Do *NOT* call _spinner function directly. 10 | # Use {start,stop}_spinner wrapper functions 11 | 12 | # usage: 13 | # 1. source this script in your's 14 | # 2. start the spinner: 15 | # start_spinner [display-message-here] 16 | # 3. run your command 17 | # 4. stop the spinner: 18 | # stop_spinner [your command's exit status] 19 | # 20 | # Also see: test.sh 21 | 22 | 23 | function _spinner() { 24 | # $1 start/stop 25 | # 26 | # on start: $2 display message 27 | # on stop : $2 process exit status 28 | # $3 spinner function pid (supplied from stop_spinner) 29 | 30 | local on_success="DONE" 31 | local on_fail="FAIL" 32 | local white="\e[1;37m" 33 | local green="\e[1;32m" 34 | local red="\e[1;31m" 35 | local nc="\e[0m" 36 | 37 | case $1 in 38 | start) 39 | # calculate the column where spinner and status msg will be displayed 40 | let column=$(tput cols)-${#2}-8 41 | # display message and position the cursor in $column column 42 | echo -ne ${2} 43 | printf "%${column}s" 44 | 45 | # start spinner 46 | i=1 47 | sp='\|/-' 48 | delay=${SPINNER_DELAY:-0.15} 49 | 50 | while : 51 | do 52 | printf "\b${sp:i++%${#sp}:1}" 53 | sleep $delay 54 | done 55 | ;; 56 | stop) 57 | if [[ -z ${3} ]]; then 58 | echo "spinner is not running.." 59 | exit 1 60 | fi 61 | 62 | kill $3 > /dev/null 2>&1 63 | 64 | # inform the user uppon success or failure 65 | echo -en "\b[" 66 | if [[ $2 -eq 0 ]]; then 67 | echo -en "${green}${on_success}${nc}" 68 | else 69 | echo -en "${red}${on_fail}${nc}" 70 | fi 71 | echo -e "]" 72 | ;; 73 | *) 74 | echo "invalid argument, try {start/stop}" 75 | exit 1 76 | ;; 77 | esac 78 | } 79 | 80 | function start_spinner { 81 | # $1 : msg to display 82 | _spinner "start" "${1}" & 83 | # set global spinner pid 84 | _sp_pid=$! 85 | disown 86 | } 87 | 88 | function stop_spinner { 89 | # $1 : command exit status 90 | _spinner "stop" $1 $_sp_pid 91 | unset _sp_pid 92 | } 93 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE TupleSections #-} 2 | module Agda.Compiler.Malfunction (backend) where 3 | 4 | import Agda.Compiler.Backend 5 | import Agda.Utils.Pretty 6 | import Control.Monad 7 | import Control.Monad.Extra 8 | import Control.Monad.Trans 9 | import Data.Bifunctor 10 | import Data.Char 11 | import Data.Either 12 | import Data.Generics.Uniplate 13 | import Data.Ix (rangeSize) 14 | import Data.Ix 15 | import Data.List 16 | import Data.Map (Map) 17 | import qualified Data.Map as Map 18 | import Data.Maybe 19 | import Data.Set (Set) 20 | import qualified Data.Set as Set 21 | import Numeric (showHex) 22 | import System.Console.GetOpt 23 | import Text.Printf 24 | 25 | import Agda.Compiler.Malfunction.AST 26 | import qualified Agda.Compiler.Malfunction.Compiler as Mlf 27 | import Agda.Compiler.Malfunction.Instances 28 | import Agda.Compiler.Malfunction.Run 29 | import qualified Agda.Compiler.Malfunction.Run as Run 30 | import Agda.Syntax.Common (NameId) 31 | 32 | backend :: Backend 33 | backend = Backend backend' 34 | 35 | data MlfOptions = Opts 36 | { _enabled :: Bool 37 | , _resultVar :: Maybe Ident 38 | , _outputFile :: Maybe FilePath 39 | , _outputMlf :: Maybe FilePath 40 | , _debug :: Bool 41 | } 42 | 43 | defOptions :: MlfOptions 44 | defOptions = Opts 45 | { _enabled = False 46 | , _resultVar = Nothing 47 | , _outputFile = Nothing 48 | , _outputMlf = Nothing 49 | , _debug = False 50 | } 51 | 52 | ttFlags :: [OptDescr (Flag MlfOptions)] 53 | ttFlags = 54 | [ Option [] ["mlf"] (NoArg $ \ o -> return o{ _enabled = True }) 55 | "Generate Malfunction" 56 | , Option ['r'] ["print-var"] (ReqArg (\r o -> return o{_resultVar = Just r}) "VAR") 57 | "(DEBUG) Run the module and print the integer value of a variable" 58 | , Option ['o'] [] (ReqArg (\r o -> return o{_outputFile = Just r}) "FILE") 59 | "(DEBUG) Place outputFile resulting module into FILE" 60 | , Option ['d'] ["debug"] (NoArg $ \ o -> return o{ _debug = True }) 61 | "Generate Malfunction" 62 | , Option [] ["compilemlf"] (ReqArg (\r o -> return o{_outputMlf = Just r}) "MODNAME") 63 | "Runs the malfunction compiler on the output file" 64 | ] 65 | 66 | backend' :: Backend' MlfOptions MlfOptions () [Definition] Definition 67 | backend' = Backend' { 68 | backendName = "malfunction" 69 | , options = defOptions 70 | , commandLineFlags = ttFlags 71 | , isEnabled = _enabled 72 | , preCompile = return 73 | , postCompile = mlfPostCompile --liftIO (putStrLn "post compile") 74 | , preModule = \_enf _m _ifile -> return $ Recompile () 75 | , compileDef = \_env _menv def -> return def 76 | , postModule = \_env _menv _m _mod defs -> return defs --mlfPostModule env defs 77 | , backendVersion = Nothing 78 | , scopeCheckingSuffices = False 79 | } 80 | 81 | definitionSummary :: MlfOptions -> Definition -> TCM () 82 | definitionSummary opts def = when (_debug opts) $ do 83 | liftIO (putStrLn ("Summary for: " ++ show q)) 84 | liftIO $ putStrLn $ unlines [ 85 | show (defName def) 86 | ++ " (" ++ show (Mlf.qnameNameId q)++ "), " ++ defntype 87 | ] 88 | case theDef def of 89 | Function{} -> 90 | whenJustM (toTreeless q) $ 91 | \tt -> 92 | liftIO . putStrLn . render 93 | $ header '=' (show q) 94 | $$ sect "Treeless (abstract syntax)" (text . show $ tt) 95 | $$ sect "Treeless (concrete syntax)" (pretty tt) 96 | _ -> return () 97 | where 98 | sect t dc = text t $+$ nest 2 dc $+$ text "" 99 | header c h = let cs = replicate 15 c 100 | in text $ printf "%s %s %s" cs h cs 101 | q = defName def 102 | defntype = case theDef def of 103 | Constructor{} -> "constructor" 104 | Primitive{} -> "primitive" 105 | Function{} -> "function" 106 | Datatype{} -> "datatype" 107 | Record{} -> "record" 108 | AbstractDefn{} -> "abstract" 109 | Axiom{} -> "axiom" 110 | 111 | -- TODO: Maybe we'd like to refactor this so that we first do something like 112 | -- this (in the calling function) 113 | -- 114 | -- partition [Definition] -> ([Function], [Primitive], ...) 115 | -- 116 | -- And then we can simply do the topologic sorting stuff on the functions only. 117 | -- This would certainly make this funciton a helluva lot cleaner. 118 | -- 119 | -- | Compiles a whole module 120 | mlfMod 121 | :: [Definition] -- ^ All visible definitions 122 | -> TCM Mod 123 | mlfMod allDefs = do 124 | grps' <- mapMaybeM act allDefs 125 | let (primsAndAxioms, tlFunBindings) = partitionEithers grps' 126 | (prims, axioms) = partitionEithers primsAndAxioms 127 | env <- getCompilerEnv (getConstructors allDefs) tlFunBindings 128 | let (MMod funBindings ts) = compile env tlFunBindings 129 | primBindings = catMaybes $ Mlf.runTranslate (mapM (uncurry Mlf.compilePrim) prims) env 130 | axiomBindings = catMaybes $ Mlf.runTranslate (mapM Mlf.compileAxiom axioms) env 131 | return $ MMod (axiomBindings ++ primBindings ++ funBindings) ts 132 | where 133 | act :: Definition -> TCM (Maybe (Either (Either (QName, String) QName) (QName, TTerm))) 134 | act def@Defn{defName = q, theDef = d} = case d of 135 | Function{} -> fmap Right <$> getBindings def 136 | Primitive{ primName = s } -> return $ Just $ Left $ Left (q, s) 137 | Axiom{} -> return $ Just $ Left $ Right q 138 | _ -> return Nothing 139 | 140 | compile :: Mlf.Env -> [(QName, TTerm)] -> Mod 141 | compile env bs = Mlf.compile env bs 142 | 143 | qnamesInTerm :: Set QName -> TTerm -> Set QName 144 | qnamesInTerm s t = go t s 145 | where 146 | go :: TTerm -> Set QName -> Set QName 147 | go t qs = case t of 148 | TDef q -> Set.insert q qs 149 | TApp f args -> foldr go qs (f:args) 150 | TLam b -> go b qs 151 | TCon q -> Set.insert q qs 152 | TLet a b -> foldr go qs [a, b] 153 | TCase _ _ p alts -> foldr qnamesInAlt (go p qs) alts 154 | _ -> qs 155 | where 156 | qnamesInAlt a qs = case a of 157 | TACon q _ t -> Set.insert q (go t qs) 158 | TAGuard t b -> foldr go qs [t, b] 159 | TALit _ b -> go b qs 160 | 161 | 162 | -- | The argument allNames is optional. If you provide an empty list concrete 163 | -- names will not be appended to the NameId 164 | getCompilerEnv :: [[QName]] -> [(QName, TTerm)] -> TCM Mlf.Env 165 | getCompilerEnv allcons bs 166 | | any ((>rangeSize Mlf.mlfTagRange) . length) allcons = error "too many constructors" 167 | | otherwise = do 168 | wa <- withArity allcons 169 | return (Mlf.mkCompilerEnv2 allNames wa) 170 | where 171 | allNames = Set.toList $ foldr step mempty bs 172 | step (qn, tt) acc = qnamesInTerm (Set.insert qn acc) tt 173 | 174 | -- | Creates a mapping for all the constructors in the array. The constructors 175 | -- should reference the same data-type. 176 | withArity :: [[QName]] -> TCM [[(QName, Int)]] 177 | withArity = mapM (mapM (\q -> (q,) <$> arityQName q)) 178 | 179 | -- | If the qnames references a constructor the arity of that constructor is returned. 180 | arityQName :: QName -> TCM Int 181 | arityQName q = f . theDef <$> getConstInfo q 182 | where 183 | f def = case def of 184 | Constructor{} -> conArity def 185 | _ -> error "Not a constructor :(" 186 | 187 | getBindings :: Definition -> TCM (Maybe (QName, TTerm)) 188 | getBindings Defn{defName = q} = fmap (\t -> (q, t)) <$> toTreeless q 189 | 190 | mlfPostCompile :: MlfOptions -> IsMain -> Map ModuleName [Definition] -> TCM () 191 | mlfPostCompile opts _ modToDefs = do 192 | mapM_ (definitionSummary opts) allDefs 193 | void $ mlfPostModule opts allDefs 194 | where 195 | allDefs :: [Definition] 196 | allDefs = concat (Map.elems modToDefs) 197 | 198 | -- TODO: `Definition`'s should be sorted *and* grouped by `defMutual` (a field 199 | -- in Definition). each group should compile to: 200 | -- 201 | -- (rec 202 | -- x0 = def0 203 | -- ... 204 | -- ) 205 | mlfPostModule :: MlfOptions -> [Definition] -> TCM Mod 206 | mlfPostModule opts defs = do 207 | modl@(MMod binds _) <- mlfMod defs 208 | let modlTxt = prettyShow $ fromMaybe modl 209 | ((withPrintInts modl . pure) <$> (_resultVar opts >>= fromSimpleIdent binds)) 210 | when (_debug opts) $ liftIO . putStrLn $ modlTxt 211 | whenJust (_resultVar opts) (printVars opts modl . pure) 212 | whenJust (_outputFile opts) (liftIO . (`writeFile`modlTxt)) 213 | whenJust (_outputMlf opts) $ \fp -> liftIO $ Run.runMalfunction fp modlTxt 214 | return modl 215 | 216 | printVars :: MonadIO m => MlfOptions -> Mod -> [Ident] -> m () 217 | printVars opts modl@(MMod binds _) simpleVars = when (_debug opts) $ do 218 | liftIO (putStrLn "\n=======================") 219 | case fullNames of 220 | Just vars -> liftIO $ runModPrintInts modl vars >>= putStrLn 221 | _ -> 222 | liftIO $ 223 | putStrLn 224 | "Variable not bound, did you specify the *fully quailified* name, e.g. \"Test.var\"?" 225 | where 226 | fullNames = mapM (fromSimpleIdent binds) simpleVars 227 | 228 | -- | "Test2.a" --> 24.1932f7ddf4cc7d3a.Test2.a 229 | fromSimpleIdent :: [Binding] -> Ident -> Maybe Ident 230 | fromSimpleIdent binds simple = listToMaybe (filter (isSuffixOf simple) (getNames binds)) 231 | where 232 | getNames = mapMaybe getName 233 | getName (Named u _) = Just u 234 | getName _ = Nothing 235 | 236 | -- | Returns all constructors grouped by data type. 237 | getConstructors :: [Definition] -> [[QName]] 238 | getConstructors = mapMaybe (getCons . theDef) 239 | where 240 | getCons :: Defn -> Maybe [QName] 241 | getCons c@Datatype{} = Just (dataCons c) 242 | -- The way I understand it a record is just like a data-type 243 | -- except it only has one constructor and that one constructor 244 | -- takes as many arguments as the number of fields in that 245 | -- record. 246 | getCons c@Record{} = Just . pure . recCon $ c 247 | -- TODO: Stub value here! 248 | getCons _ = Nothing 249 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/AST.hs: -------------------------------------------------------------------------------- 1 | {- | 2 | Module : Agda.Compiler.Malfunction.AST 3 | Maintainer : janmasrovira@gmail.com, hanghj@student.chalmers.se 4 | 5 | This module defines the abstract syntax of 6 | . Please see the 7 | 9 | -} 10 | {-# LANGUAGE OverloadedStrings #-} 11 | module Agda.Compiler.Malfunction.AST 12 | ( IntType(..) 13 | , IntConst(..) 14 | , UnaryIntOp(..) 15 | , BinaryIntOp(..) 16 | , VectorType(..) 17 | , Mutability(..) 18 | , BlockTag 19 | , Case(..) 20 | , Ident 21 | , Longident 22 | , Mod(..) 23 | , Term(..) 24 | , Binding(..) 25 | -- NOTE: I don't know which of these is preferable 26 | -- * Don't re-export anything from Agda.Utils.Pretty 27 | -- * export a few things (like we do currently) 28 | -- * Re-export the whole module 29 | , pretty 30 | , prettyShow 31 | ) where 32 | 33 | import Data.Int 34 | -- There does exist a definition of a type-class `Pretty` in the package 35 | -- `pretty` but this is not the one used for Treeless terms, so for consistency, 36 | -- let's go with Agda's choice. 37 | import Agda.Utils.Pretty 38 | 39 | -- | The integer types. 40 | data IntType 41 | = TInt 42 | | TInt32 43 | | TInt64 44 | | TBigint 45 | deriving (Show, Eq) 46 | 47 | -- | An integer value tagged with its corresponding type. 48 | data IntConst 49 | -- Int may be the wrong type. 50 | -- 51 | -- In malfunction Int is: 52 | -- 53 | -- > `int` uses either 31- or 63- bit two's complement arithmetic 54 | -- > (depending on system word size, and also wrapping on overflow) 55 | -- > https://github.com/stedolan/malfunction/blob/master/docs/spec.md 56 | -- 57 | -- And in Haskell: 58 | -- > A fixed-precision integer type with at least the range 59 | -- > [-2^29 .. 2^29-1] 60 | -- > https://hackage.haskell.org/package/base-4.9.0.0/docs/Data-Int.html 61 | -- 62 | -- Jan: Just run 63 | -- logBase 2 ((fromIntegral (maxBound :: Int)) :: Double) 64 | -- in my computer (64 bits) --> 2 ^ 63 - 1 == (maxBound :: Int) 65 | = CInt Int 66 | | CInt32 Int32 67 | | CInt64 Int64 68 | | CBigint Integer 69 | deriving (Show, Eq) 70 | 71 | -- | A unary operator. 72 | data UnaryIntOp = Neg | Not 73 | deriving (Show, Eq) 74 | 75 | -- | A binary operator. 76 | data BinaryIntOp 77 | = Add | Sub | Mul | Div | Mod | And | Or | Xor | Lsl | Lsr | Asr 78 | | Lt | Gt | Lte | Gte | Eq 79 | deriving (Show, Eq) 80 | 81 | -- | Vector types. 82 | data VectorType = Array | Bytevec 83 | deriving (Show, Eq) 84 | 85 | -- | Mutability 86 | data Mutability = Inm | Mut 87 | deriving (Show, Eq) 88 | 89 | -- NOTE: Any tag value above 200 is an error in malfunction. 90 | -- 91 | -- For this reason we may want to make BlockTag a newtype and only export a constructor. 92 | -- 93 | -- | A tag used in the construction of $Block@s. 94 | type BlockTag = Int 95 | 96 | -- The spec and the ocaml implementation are inconsistent when defining Case. 97 | -- I'll use the definition (examples) from the spec to guide this implementation. 98 | -- I know I could've used Maybe's here, but not doing so was a concious choice. 99 | -- 100 | -- | Used in switch-statements. See 101 | -- 102 | data Case 103 | -- (tag _) 104 | = Deftag 105 | -- (tag n) 106 | | Tag Int 107 | -- _ 108 | | CaseAnyInt 109 | -- n 110 | | CaseInt Int 111 | -- (n m) 112 | | Intrange (Int, Int) 113 | deriving (Show, Eq) 114 | 115 | -- | An identifier used to reference other values in the malfunction module. 116 | type Ident = String 117 | 118 | -- | A long identifier is used to reference OCaml values (using @Mglobal@). 119 | type Longident = [Ident] 120 | 121 | --data Longident 122 | -- = Lident String 123 | -- | Ldot Longident String 124 | -- | Lapply Longident Longident 125 | 126 | -- "The top-level sexp must begin with the atom module, followed by a 127 | -- list of bindings (described under let, below), followed by an sexp 128 | -- beginning with the atom export." 129 | -- 130 | -- | Defines a malfunction module. 131 | data Mod = MMod [Binding] [Term] deriving (Eq, Show) 132 | 133 | -- | The overall syntax of malfunction terms. 134 | data Term 135 | = Mvar Ident 136 | | Mlambda [Ident] Term 137 | | Mapply Term [Term] 138 | | Mlet [Binding] Term 139 | | Mint IntConst 140 | | Mstring String 141 | | Mglobal Longident 142 | | Mswitch Term [([Case], Term)] 143 | -- Integers 144 | | Mintop1 UnaryIntOp IntType Term 145 | | Mintop2 BinaryIntOp IntType Term Term 146 | | Mconvert IntType IntType Term 147 | -- Vectors 148 | | Mvecnew VectorType Term Term 149 | | Mvecget VectorType Term Term 150 | | Mvecset VectorType Term Term Term 151 | | Mveclen VectorType Term 152 | -- Blocks 153 | | Mblock Int [Term] 154 | | Mfield Int Term 155 | deriving (Show, Eq) 156 | 157 | -- | Bindings 158 | -- 159 | -- The atom `let` introduces a sequence of bindings: 160 | -- 161 | -- (let BINDING BINDING BINDING ... BODY) 162 | -- 163 | -- Each binding is of one of the forms: 164 | -- 165 | -- - @Named@: @($var EXP)@: binds @$var@ to the result of evaluating @EXP@. 166 | -- @$var@ scopes over subsequent bindings and the body. 167 | -- 168 | -- - @Unnamed@: @(_ EXP)@: evaluates @EXP@ and ignores the result. 169 | -- 170 | -- - @Recursive@: @(rec ($VAR1 EXP1) ($VAR2 EXP2) ...)@: binds each @$VAR@ mutually 171 | -- recursively. Each @EXP@ must be of the form @(lambda ...)@. 172 | -- Bindings scope over themselves, each other, subsequent 173 | -- bindings, and the body. 174 | data Binding 175 | = Unnamed Term 176 | | Named Ident Term 177 | | Recursive [(Ident, Term)] 178 | deriving (Show, Eq) 179 | 180 | textShow :: Show a => a -> Doc 181 | textShow = text . show 182 | 183 | nst :: Doc -> Doc 184 | nst = nest 2 185 | 186 | (<.>) :: Doc -> Doc -> Doc 187 | a <.> b = a <> "." <> b 188 | 189 | level :: Doc -> Doc -> Doc 190 | level a b = sep [ "(" <+> a, nst b, ")" ] 191 | 192 | levelPlus :: Doc -> [Doc] -> Doc 193 | levelPlus a bs = sep $ [ "(" <+> a ] ++ map nst bs ++ [")"] 194 | 195 | instance Pretty Mod where 196 | pretty (MMod bs ts) = levelPlus "module" (map pretty bs ++ [levelPlus "export" (map pretty ts)]) 197 | prettyPrec _ = pretty 198 | 199 | instance Pretty Term where 200 | pretty tt = case tt of 201 | Mvar i -> prettyIdent i 202 | Mlambda is t -> level ("lambda" <+> parens (hsep (map prettyIdent is))) (pretty t) 203 | Mapply t ts -> levelPlus ("apply " <> pretty t) (map pretty ts) 204 | Mlet bs t -> level "let" (prettyList bs $$ pretty t) 205 | Mint ic -> pretty ic 206 | Mstring s -> textShow s 207 | Mglobal li -> parens $ "global" <+> prettyLongident li 208 | Mswitch t cexps -> levelPlus ("switch" <+> pretty t) (map prettyCaseExpression cexps) 209 | -- Integers 210 | Mintop1 op tp t0 -> pretty op <+> prettyTypedTerm tp t0 211 | Mintop2 op tp t0 t1 -> levelPlus (pretty op) [prettyTypedTerm tp t0, prettyTypedTerm tp t1] 212 | Mconvert tp0 tp1 t0 -> parens $ "convert" <.> pretty tp0 <.> pretty tp1 <+> pretty t0 213 | -- Vectors 214 | Mvecnew _tp t0 t1 -> levelPlus "makevec" [pretty t0, pretty t1] 215 | Mvecget _tp t0 t1 -> levelPlus "load" [pretty t0, pretty t1] 216 | Mvecset _tp t0 t1 t2 -> levelPlus "store" [pretty t0, pretty t1, pretty t2] 217 | Mveclen _tp t0 -> level "length" (pretty t0) 218 | -- Blocks 219 | Mblock i ts -> level ("block" <+> parens ("tag" <+> pretty i)) (prettyList ts) 220 | Mfield i t0 -> parens $ "field" <+> pretty i <+> pretty t0 221 | prettyPrec _ = pretty 222 | 223 | instance Pretty Binding where 224 | pretty b = case b of 225 | Unnamed t -> level "_" (pretty t) 226 | Named i t -> level (prettyIdent i) (pretty t) 227 | Recursive bs -> levelPlus "rec" (map showIdentTerm bs) 228 | where 229 | showIdentTerm :: (Ident, Term) -> Doc 230 | showIdentTerm (i, t) = level (prettyIdent i) (pretty t) 231 | 232 | instance Pretty IntConst where 233 | pretty ic = case ic of 234 | CInt i -> pretty i 235 | CInt32 i -> pretty i 236 | CInt64 i -> textShow i 237 | CBigint i -> pretty i 238 | 239 | prettyLongident :: Longident -> Doc 240 | prettyLongident = hsep . map prettyIdent 241 | 242 | prettyIdent :: Ident -> Doc 243 | prettyIdent = text . ('$':) 244 | 245 | prettyCaseExpression :: ([Case], Term) -> Doc 246 | prettyCaseExpression (cs, t) = level (prettyList cs) (pretty t) 247 | 248 | instance Pretty Case where 249 | pretty c = case c of 250 | Deftag -> "(tag _)" 251 | Tag n -> "(tag " <> pretty n <> ")" 252 | CaseAnyInt -> "_" 253 | CaseInt n -> pretty n 254 | Intrange (i, j) -> "(" <> pretty i <+> pretty j <> ")" 255 | 256 | instance Pretty UnaryIntOp where 257 | pretty op = case op of 258 | Neg -> "?" 259 | Not -> "?" 260 | 261 | instance Pretty BinaryIntOp where 262 | pretty op = case op of 263 | Add -> "+" 264 | Sub -> "-" 265 | Mul -> "*" 266 | Div -> "/" 267 | Mod -> "%" 268 | And -> "&" 269 | Or -> "|" 270 | Xor -> "^" 271 | Lsl -> "<<" 272 | Lsr -> ">>" 273 | Asr -> "a>>" 274 | Lt -> "<" 275 | Gt -> ">" 276 | Lte -> "<=" 277 | Gte -> ">=" 278 | Eq -> "==" 279 | 280 | prettyTypedTerm :: IntType -> Term -> Doc 281 | prettyTypedTerm tp t = case tp of 282 | TInt -> pretty t 283 | _ -> pretty t <.> pretty tp 284 | 285 | instance Pretty IntType where 286 | pretty tp = case tp of 287 | TInt -> "int" 288 | TInt32 -> "int32" 289 | TInt64 -> "int64" 290 | TBigint -> "bigint" 291 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/Compiler.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE RecordWildCards #-} 2 | {- | 3 | Module : Agda.Compiler.Malfunction.Compiler 4 | Maintainer : janmasrovira@gmail.com, hanghj@student.chalmers.se 5 | 6 | This module includes functions that compile from to 7 | . 8 | -} 9 | {-# LANGUAGE FlexibleContexts #-} 10 | {-# LANGUAGE ConstraintKinds #-} 11 | module Agda.Compiler.Malfunction.Compiler 12 | ( 13 | -- * Translation functions 14 | translateTerms 15 | , translateDef 16 | , nameToIdent 17 | , compile 18 | , runTranslate 19 | -- * Data needed for compilation 20 | , Env(..) 21 | , ConRep(..) 22 | , Arity 23 | -- , mkCompilerEnv 24 | , mkCompilerEnv2 25 | -- * Others 26 | , qnameNameId 27 | , errorT 28 | , boolT 29 | , wildcardTerm 30 | , namedBinding 31 | , nameIdToIdent 32 | , nameIdToIdent' 33 | , mlfTagRange 34 | -- * Primitives 35 | , compilePrim 36 | , compileAxiom 37 | -- * Malfunction AST 38 | , module Agda.Compiler.Malfunction.AST 39 | ) where 40 | 41 | import Agda.Syntax.Common (NameId(..)) 42 | import Agda.Syntax.Literal 43 | import Agda.Syntax.Treeless 44 | 45 | import Control.Monad 46 | import Control.Monad.Extra 47 | import Control.Monad.Identity 48 | import Data.List.Extra 49 | import Control.Monad.Reader 50 | import Data.Graph 51 | import Data.Ix 52 | import Data.List 53 | import Data.Map (Map) 54 | import qualified Data.Map as Map 55 | import Data.Maybe 56 | import Data.Set (Set) 57 | import qualified Data.Set as Set 58 | import Data.Tuple.Extra 59 | import Numeric (showHex) 60 | import Data.Char 61 | 62 | import Agda.Compiler.Malfunction.AST 63 | import qualified Agda.Compiler.Malfunction.Primitive as Primitive 64 | 65 | data Env = Env 66 | { _conMap :: Map NameId ConRep 67 | , _qnameConcreteMap :: Map NameId String 68 | , _level :: Int 69 | , _biBool :: Maybe (NameId, NameId) 70 | } 71 | deriving (Show) 72 | 73 | -- | Data needed to represent a constructor 74 | data ConRep = ConRep 75 | { _conTag :: Int 76 | , _conArity :: Int 77 | } deriving (Show) 78 | 79 | type Translate a = Reader Env a 80 | type Arity = Int 81 | 82 | runTranslate :: Reader Env a -> Env -> a 83 | runTranslate = runReader 84 | 85 | translateDefM :: MonadReader Env m => QName -> TTerm -> m Binding 86 | translateDefM qnm t 87 | | isRecursive = do 88 | tt <- translateM t 89 | iden <- nameToIdent qnm 90 | return . Recursive . pure $ (iden, tt) 91 | | otherwise = do 92 | tt <- translateM t 93 | namedBinding qnm tt 94 | where 95 | -- TODO: I don't believe this is enough, consider the example 96 | -- where functions are mutually recursive. 97 | -- a = b 98 | -- b = a 99 | isRecursive = Set.member (qnameNameId qnm) (qnamesIdsInTerm t) -- TODO: is this enough? 100 | 101 | mkCompilerEnv :: [QName] -> Map NameId ConRep -> Env 102 | mkCompilerEnv allNames conMap = Env { 103 | _conMap = conMap 104 | , _level = 0 105 | , _qnameConcreteMap = qnameMap 106 | , _biBool = Nothing 107 | } 108 | where 109 | qnameMap = Map.fromList [ (qnameNameId qn, concreteName qn) | qn <- allNames ] 110 | showNames = intercalate "." . map (concatMap toValid . show . nameConcrete) 111 | concreteName qn = showNames (mnameToList (qnameModule qn) ++ [qnameName qn]) 112 | toValid :: Char -> String 113 | toValid c 114 | | any (`inRange`c) [('0','9'), ('a', 'z'), ('A', 'Z')] 115 | || c`elem`"_" = [c] 116 | | otherwise = "{" ++ show (ord c) ++ "}" 117 | 118 | mlfTagRange :: (Int, Int) 119 | mlfTagRange = (0, 199) 120 | 121 | mkCompilerEnv2 :: [QName] -> [[(QName, Arity)]] -> Env 122 | mkCompilerEnv2 allNames consByDtype = Env { 123 | _conMap = conMap 124 | , _level = 0 125 | , _qnameConcreteMap = qnameMap 126 | , _biBool = findBuiltinBool (map (map fst) consByDtype) 127 | } 128 | where 129 | conMap = Map.fromList [ (qnameNameId qn, ConRep {..} ) 130 | | typeCons <- consByDtype 131 | , (length consByDtype <= rangeSize mlfTagRange) 132 | || (error "too many constructors") 133 | , (_conTag, (qn, _conArity)) <- zip (range mlfTagRange) typeCons ] 134 | qnameMap = Map.fromList [ (qnameNameId qn, concreteName qn) | qn <- allNames ] 135 | showNames = intercalate "." . map (concatMap toValid . show . nameConcrete) 136 | concreteName qn = showNames (mnameToList (qnameModule qn) ++ [qnameName qn]) 137 | toValid :: Char -> String 138 | toValid c 139 | | any (`inRange`c) [('0','9'), ('a', 'z'), ('A', 'Z')] 140 | || c`elem`"_" = [c] 141 | | otherwise = "{" ++ show (ord c) ++ "}" 142 | 143 | -- | Translate a single treeless term to a list of malfunction terms. 144 | -- 145 | -- Note that this does not handle mutual dependencies correctly. For this you 146 | -- would need @compile@. 147 | translateDef :: Env -> QName -> TTerm -> Binding 148 | translateDef env qn = (`runTranslate` env) . translateDefM qn 149 | 150 | -- | Translates a list treeless terms to a list of malfunction terms. 151 | -- 152 | -- Pluralized version of @translateDef@. 153 | translateTerms :: Env -> [TTerm] -> [Term] 154 | translateTerms env = (`runTranslate` env) . mapM translateM 155 | 156 | translateM :: MonadReader Env m => TTerm -> m Term 157 | translateM = translateTerm 158 | 159 | translateTerm :: MonadReader Env m => TTerm -> m Term 160 | translateTerm tt = case tt of 161 | TVar i -> indexToVarTerm i 162 | TPrim tp -> return $ translatePrimApp tp [] 163 | TDef name -> translateName name 164 | TApp t0 args -> translateApp t0 args 165 | TLam{} -> translateLam tt 166 | TLit lit -> return $ translateLit lit 167 | TCon nm -> translateCon nm [] 168 | TLet t0 t1 -> do 169 | t0' <- translateTerm t0 170 | (var, t1') <- introVar (translateTerm t1) 171 | return (Mlet [Named var t0'] t1') 172 | -- @deflt@ is the default value if all @alt@s fail. 173 | TCase i _ deflt alts -> do 174 | t <- indexToVarTerm i 175 | alts' <- alternatives t 176 | return $ Mswitch t alts' 177 | where 178 | -- Case expressions may not have an alternative, this is encoded 179 | -- by @deflt@ being TError TUnreachable. 180 | alternatives t = case deflt of 181 | TError TUnreachable -> translateAltsChain t Nothing alts 182 | _ -> do 183 | d <- translateTerm deflt 184 | translateAltsChain t (Just d) alts 185 | TUnit -> return unitT 186 | TSort -> error ("Unimplemented " ++ show tt) 187 | TErased -> return wildcardTerm -- TODO: so... anything can go here? 188 | TError TUnreachable -> return wildcardTerm 189 | 190 | -- | We use this when we don't care about the translation. 191 | wildcardTerm :: Term 192 | wildcardTerm = nullary $ errorT "__UNREACHABLE__" 193 | 194 | nullary :: Term -> Term 195 | nullary = Mlambda [] 196 | 197 | indexToVarTerm :: MonadReader Env m => Int -> m Term 198 | indexToVarTerm i = do 199 | ni <- asks _level 200 | return (Mvar (ident (ni - i - 1))) 201 | 202 | -- translateSwitch :: MonadReader m => Term -> TAlt -> m ([Case], Term) 203 | -- translateSwitch tcase alt = case alt of 204 | -- -- TAGuard c t -> liftM2 (,) (pure <$> translateCase c) (translateTerm t) 205 | -- TALit pat body -> do 206 | -- b <- translateTerm body 207 | -- let c = pure $ litToCase pat 208 | -- return (c, b) 209 | -- TACon con arity t -> do 210 | -- tg <- nameToTag con 211 | -- usedFields <- snd <$> introVars arity 212 | -- (Set.map (\ix -> arity - ix - 1) . Set.filter ( usedVars t) 213 | -- (vars, t') <- introVars arity (translateTerm t) 214 | -- let bt = bindFields vars usedFields tcase t' 215 | -- -- TODO: It is not clear how to deal with bindings in a pattern 216 | -- return (pure tg, bt) 217 | -- TAGuard gd rhs -> return ([], Mvar "TAGuard.undefined") 218 | 219 | translateAltsChain :: MonadReader Env m => Term -> Maybe Term -> [TAlt] -> m [([Case], Term)] 220 | translateAltsChain tcase defaultt [] = return $ maybe [] (\d -> [(defaultCase, d)]) defaultt 221 | translateAltsChain tcase defaultt (ta:tas) = 222 | case ta of 223 | TALit pat body -> do 224 | b <- translateTerm body 225 | let c = litToCase pat 226 | (([c], b):) <$> go 227 | TACon con arity t -> do 228 | tg <- nameToTag con 229 | usedFields <- snd <$> introVars arity 230 | (Set.map (\ix -> arity - ix - 1) . Set.filter ( usedVars t) 231 | (vars, t') <- introVars arity (translateTerm t) 232 | let bt = bindFields vars usedFields tcase t' 233 | -- TODO: It is not clear how to deal with bindings in a pattern 234 | (([tg], bt):) <$> go 235 | TAGuard grd t -> do 236 | tgrd <- translateTerm grd 237 | t' <- translateTerm t 238 | tailAlts <- go 239 | return [(defaultCase, 240 | Mswitch tgrd 241 | [(trueCase, t') 242 | , (defaultCase, Mswitch tcase tailAlts)])] 243 | where 244 | go = translateAltsChain tcase defaultt tas 245 | 246 | defaultCase :: [Case] 247 | defaultCase = [CaseAnyInt, Deftag] 248 | 249 | bindFields :: [Ident] -> Set Int -> Term -> Term -> Term 250 | bindFields vars used termc body = case map bind varsRev of 251 | [] -> body 252 | binds -> Mlet binds body 253 | where 254 | varsRev = zip [0..] (reverse vars) 255 | arity = length vars 256 | bind (ix, iden) 257 | -- TODO: we bind all fields. The detection of used fields is bugged. 258 | | True || Set.member ix used = Named iden (Mfield ix termc) 259 | | otherwise = Named iden wildcardTerm 260 | 261 | litToCase :: Literal -> Case 262 | litToCase l = case l of 263 | LitNat _ i -> CaseInt . fromInteger $ i 264 | _ -> error "Unimplemented" 265 | 266 | -- The argument is the lambda itself and not its body. 267 | translateLam :: MonadReader Env m => TTerm -> m Term 268 | translateLam lam = do 269 | (is, t) <- translateLams lam 270 | return $ Mlambda is t 271 | where 272 | translateLams :: MonadReader Env m => TTerm -> m ([Ident], Term) 273 | translateLams (TLam body) = do 274 | (thisVar, (xs, t)) <- introVar (translateLams body) 275 | return (thisVar:xs, t) 276 | translateLams e = do 277 | e' <- translateTerm e 278 | return ([], e') 279 | 280 | introVars :: MonadReader Env m => Int -> m a -> m ([Ident], a) 281 | introVars k ma = do 282 | (names, env') <- nextIdxs k 283 | r <- local (const env') ma 284 | return (names, r) 285 | where 286 | nextIdxs :: MonadReader Env m => Int -> m ([Ident], Env) 287 | nextIdxs k = do 288 | i0 <- asks _level 289 | e <- ask 290 | return (map ident $ reverse [i0..i0 + k - 1], e{_level = _level e + k}) 291 | 292 | introVar :: MonadReader Env m => m a -> m (Ident, a) 293 | introVar ma = first head <$> introVars 1 ma 294 | 295 | -- This is really ugly, but I've done this for the reason mentioned 296 | -- in `translatePrim'`. Note that a similiar "optimization" could be 297 | -- done for chained lambda-expressions: 298 | -- TLam (TLam ...) 299 | translateApp :: MonadReader Env m => TTerm -> [TTerm] -> m Term 300 | translateApp ft xst = case ft of 301 | TPrim p -> translatePrimApp p <$> mapM translateTerm xst 302 | TCon nm -> translateCon nm xst 303 | _ -> do 304 | f <- translateTerm ft 305 | xs <- mapM translateTerm xst 306 | return $ Mapply f xs 307 | 308 | ident :: Int -> Ident 309 | ident i = "v" ++ show i 310 | 311 | translateLit :: Literal -> Term 312 | translateLit l = case l of 313 | LitNat _ x -> Mint (CBigint x) 314 | LitString _ s -> Mstring s 315 | LitChar _ c-> Mint . CInt . fromEnum $ c 316 | _ -> error "unsupported literal type" 317 | 318 | translatePrimApp :: TPrim -> [Term] -> Term 319 | translatePrimApp tp args = 320 | case tp of 321 | PAdd -> intbinop Add 322 | PSub -> intbinop Sub 323 | PMul -> intbinop Mul 324 | PQuot -> intbinop Div 325 | PRem -> intbinop Mod 326 | PGeq -> intbinop Gte 327 | PLt -> intbinop Lt 328 | PEqI -> intbinop Eq 329 | PEqF -> wrong 330 | PEqS -> wrong 331 | PEqC -> intbinop Eq 332 | PEqQ -> wrong 333 | PIf -> wrong 334 | PSeq -> pseq 335 | where 336 | aType = TInt 337 | intbinop op = case args of 338 | [a, b] -> Mintop2 op aType a b 339 | [a] -> Mlambda ["b"] $ Mintop2 op aType a (Mvar "b") 340 | [] -> Mlambda ["a", "b"] $ Mintop2 op aType (Mvar "a") (Mvar "b") 341 | _ -> wrongargs 342 | -- NOTE: pseq is simply (\a b -> b) because malfunction is a strict language 343 | pseq = case args of 344 | [_, b] -> b 345 | [_] -> Mlambda ["b"] $ Mvar "b" 346 | [] -> Mlambda ["a", "b"] $ Mvar "b" 347 | _ -> wrongargs 348 | -- TODO: Stub! 349 | -- wrong = return $ errorT $ "stub : " ++ show tp 350 | wrongargs = error "unexpected number of arguments" 351 | wrong = undefined 352 | 353 | 354 | -- FIXME: Please not the multitude of interpreting QName in the following 355 | -- section. This may be a problem. 356 | -- This is due to the fact that QName can refer to constructors and regular 357 | -- bindings, I think we want to handle these two cases separately. 358 | 359 | -- Questionable implementation: 360 | nameToTag :: MonadReader Env m => QName -> m Case 361 | nameToTag nm = do 362 | e <- ask 363 | builtinbool <- builtinBool (qnameNameId nm) 364 | case builtinbool of 365 | Just b -> return (CaseInt (boolToInt b)) 366 | Nothing -> 367 | ifM (isConstructor nm) 368 | (Tag <$> askConTag nm) 369 | (error $ "nameToTag only implemented for constructors, qname=" ++ show nm 370 | ++ "\nenv:" ++ show e) 371 | -- (return . Tag . fromEnum . nameId . qnameName $ nm) 372 | 373 | 374 | isConstructor :: MonadReader Env m => QName -> m Bool 375 | isConstructor nm = (qnameNameId nm`Map.member`) <$> askConMap 376 | 377 | askConMap :: MonadReader Env m => m (Map NameId ConRep) 378 | askConMap = asks _conMap 379 | 380 | -- | 381 | -- Set of indices of the variables that are referenced inside the term. 382 | -- 383 | -- Example 384 | -- λλ Env{_level = 2} usedVars (λ(λ ((Var 3) (λ (Var 4)))) ) == {1} 385 | usedVars :: MonadReader Env m => TTerm -> m (Set Int) 386 | usedVars term = asks _level >>= go mempty term 387 | where 388 | go vars t topnext = goterm vars t 389 | where 390 | goterms vars = foldM (\acvars tt -> goterm acvars tt) vars 391 | goterm vars t = do 392 | nextix <- asks _level 393 | case t of 394 | (TVar v) -> return $ govar vars v nextix 395 | (TApp t args) -> goterms vars (t:args) 396 | (TLam t) -> snd <$> introVar (goterm vars t) 397 | (TLet t1 t2) -> do 398 | vars1 <- goterm vars t1 399 | snd <$> introVar (goterm vars1 t2) 400 | (TCase v _ def alts) -> do 401 | vars1 <- goterm (govar vars v nextix) def 402 | foldM (\acvars alt -> goalt acvars alt) vars1 alts 403 | _ -> return vars 404 | govar vars v nextix 405 | | 0 <= v' && v' < topnext = Set.insert v' vars 406 | | otherwise = vars 407 | where v' = v + topnext - nextix 408 | goalt vars alt = case alt of 409 | TACon _ _ b -> goterm vars b 410 | TAGuard g b -> goterms vars [g, b] 411 | TALit{} -> return vars 412 | 413 | 414 | -- TODO: Translate constructors differently from names. 415 | -- Don't know if we should do the same when translating TDef's, but here we 416 | -- should most likely use malfunction "blocks" to represent constructors 417 | -- in an "untyped but injective way". That is, we only care that each 418 | -- constructor maps to a unique number such that we will be able to 419 | -- distinguish it in malfunction. This also means that we should carry 420 | -- some state around mapping each constructor to it's corresponding 421 | -- "block-representation". 422 | -- 423 | -- An example for clarity. Consider type: 424 | -- 425 | -- T a b = L a | R b | B a b | E 426 | -- 427 | -- We need to encode the constructors in an injective way and we need to 428 | -- encode the arity of the constructors as well. 429 | -- 430 | -- translate (L a) = (block (tag 2) (tag 0) a') 431 | -- translate (R b) = (block (tag 2) (tag 1) b') 432 | -- translate (B a b) = (block (tag 3) (tag 2) a' b') 433 | -- translate E = (block (tag 1) (tag 3)) 434 | -- TODO: If the length of `ts` does not match the arity of `nm` then a lambda-expression must be returned. 435 | translateCon :: MonadReader Env m => QName -> [TTerm] -> m Term 436 | translateCon nm ts = do 437 | builtinbool <- builtinBool (qnameNameId nm) 438 | case builtinbool of 439 | Just t -> return (boolT t) 440 | Nothing -> do 441 | ts' <- mapM translateTerm ts 442 | tag <- askConTag nm 443 | arity <- askArity nm 444 | let diff = arity - length ts' 445 | vs = take diff $ map pure ['a'..] 446 | return $ if diff == 0 447 | then Mblock tag ts' 448 | else Mlambda vs (Mblock tag (ts' ++ map Mvar vs)) 449 | 450 | -- | Ugly hack to represent builtin bools as integers. 451 | -- For now it checks whether the concrete name ends with "Bool.true" or "Bool.false" 452 | builtinBool :: MonadReader Env m => NameId -> m (Maybe Bool) 453 | builtinBool qn = do 454 | isTrue <- isBuiltinTrue qn 455 | if isTrue then return (Just True) 456 | else do 457 | isFalse <- isBuiltinFalse qn 458 | if isFalse then return (Just False) 459 | else return Nothing 460 | where 461 | isBuiltinTrue :: MonadReader Env m => NameId -> m Bool 462 | isBuiltinTrue qn = maybe False ((==qn) . snd) <$> asks _biBool 463 | isBuiltinFalse :: MonadReader Env m => NameId -> m Bool 464 | isBuiltinFalse qn = maybe False ((==qn) . fst) <$> asks _biBool 465 | 466 | -- | The argument are all data constructors grouped by datatype. 467 | -- returns Maybe (false NameId, true NameId) 468 | findBuiltinBool :: [[QName]] -> Maybe (NameId, NameId) 469 | findBuiltinBool = firstJust maybeBool 470 | where maybeBool l@[_,_] = firstJust falseTrue (permutations l) 471 | where falseTrue [f, t] 472 | | "Bool.false" `isSuffixOf` show f 473 | && "Bool.true" `isSuffixOf` show t = Just (qnameNameId f, qnameNameId t) 474 | falseTrue _ = Nothing 475 | maybeBool _ = Nothing 476 | 477 | askArity :: MonadReader Env m => QName -> m Int 478 | askArity = fmap _conArity . nontotalLookupConRep 479 | 480 | askConTag :: MonadReader Env m => QName -> m Int 481 | askConTag = fmap _conTag . nontotalLookupConRep 482 | 483 | nontotalLookupConRep :: MonadReader Env f => QName -> f ConRep 484 | nontotalLookupConRep q = fromMaybe err <$> lookupConRep q 485 | where 486 | err = error $ "Could not find constructor with qname: " ++ show q 487 | 488 | lookupConRep :: MonadReader Env f => QName -> f (Maybe ConRep) 489 | lookupConRep ns = Map.lookup (qnameNameId ns) <$> asks _conMap 490 | 491 | -- Unit is treated as a glorified value in Treeless, luckily it's fairly 492 | -- straight-forward to encode using the scheme described in the documentation 493 | -- for `translateCon`. 494 | unitT :: Term 495 | unitT = Mblock 0 [] 496 | 497 | translateName :: MonadReader Env m => QName -> m Term 498 | translateName qn = Mvar <$> nameToIdent qn 499 | 500 | -- | Translate a Treeless name to a valid identifier in Malfunction 501 | -- 502 | -- Not all names in agda are valid names in Treleess. Valid names in Agda are 503 | -- given by [1]. Valid identifiers in Malfunction is subject to change: 504 | -- 505 | -- "Atoms: sequences of ASCII letters, digits, or symbols (the exact set of 506 | -- allowed symbols isn't quite nailed down yet)"[2] 507 | -- 508 | -- [1. The Agda Wiki]: 509 | -- [2. Malfunction Spec]: 510 | nameToIdent :: MonadReader Env m => QName -> m Ident 511 | nameToIdent qn = nameIdToIdent (qnameNameId qn) 512 | 513 | nameIdToIdent' :: NameId -> Maybe String -> Ident 514 | nameIdToIdent' (NameId a b) msuffix = (hex a ++ "." ++ hex b ++ suffix) 515 | where 516 | suffix = maybe "" ('.':) msuffix 517 | hex = (`showHex` "") . toInteger 518 | 519 | nameIdToIdent :: MonadReader Env m => NameId -> m Ident 520 | nameIdToIdent nid = do 521 | x <- Map.lookup nid <$> asks _qnameConcreteMap 522 | return (nameIdToIdent' nid x) 523 | 524 | -- | Translates a treeless identifier to a malfunction identifier. 525 | qnameNameId :: QName -> NameId 526 | qnameNameId = nameId . qnameName 527 | 528 | -- | Compiles treeless "bindings" to a malfunction module given groups of defintions. 529 | compile 530 | :: Env -- ^ Environment. 531 | -> [(QName, TTerm)] -- ^ List of treeless bindings. 532 | -> Mod 533 | compile env bs = runTranslate (compileM bs) env 534 | 535 | 536 | compileM :: MonadReader Env m => [(QName, TTerm)] -> m Mod 537 | compileM allDefs = do 538 | bs <- mapM translateSCC recGrps 539 | return $ MMod bs [] 540 | where 541 | translateSCC scc = case scc of 542 | AcyclicSCC single -> uncurry translateBinding single 543 | CyclicSCC grp -> translateMutualGroup grp 544 | recGrps :: [SCC (QName, TTerm)] 545 | recGrps = dependencyGraph allDefs 546 | 547 | translateMutualGroup :: MonadReader Env m => [(QName, TTerm)] -> m Binding 548 | translateMutualGroup bs = Recursive <$> mapM (uncurry translateBindingPair) bs 549 | 550 | translateBinding :: MonadReader Env m => QName -> TTerm -> m Binding 551 | translateBinding q t = uncurry Named <$> translateBindingPair q t 552 | 553 | translateBindingPair :: MonadReader Env m => QName -> TTerm -> m (Ident, Term) 554 | translateBindingPair q t = do 555 | iden <- nameToIdent q 556 | (\t' -> (iden, t')) <$> translateTerm t 557 | 558 | dependencyGraph :: [(QName, TTerm)] -> [SCC (QName, TTerm)] 559 | dependencyGraph qs = stronglyConnComp [ ((qn, tt), qnameNameId qn, edgesFrom tt) 560 | | (qn, tt) <- qs ] 561 | where edgesFrom = Set.toList . qnamesIdsInTerm 562 | 563 | 564 | qnamesIdsInTerm :: TTerm -> Set NameId 565 | qnamesIdsInTerm t = go t mempty 566 | where 567 | insertId q = Set.insert (qnameNameId q) 568 | go :: TTerm -> Set NameId -> Set NameId 569 | go t qs = case t of 570 | TDef q -> insertId q qs 571 | TApp f args -> foldr go qs (f:args) 572 | TLam b -> go b qs 573 | TCon q -> insertId q qs 574 | TLet a b -> foldr go qs [a, b] 575 | TCase _ _ p alts -> foldr qnamesInAlt (go p qs) alts 576 | _ -> qs 577 | where 578 | qnamesInAlt a qs = case a of 579 | TACon q _ t -> insertId q (go t qs) 580 | TAGuard t b -> foldr go qs [t, b] 581 | TALit _ b -> go b qs 582 | 583 | -- | Defines a run-time error in Malfunction - equivalent to @error@ in Haskell. 584 | errorT :: String -> Term 585 | errorT err = Mapply (Mglobal ["Pervasives", "failwith"]) [Mstring err] 586 | 587 | -- | Encodes a boolean value as a numerical Malfunction value. 588 | boolT :: Bool -> Term 589 | boolT b = Mint (CInt $ boolToInt b) 590 | 591 | boolToInt :: Bool -> Int 592 | boolToInt b = if b then 1 else 0 593 | 594 | trueCase :: [Case] 595 | trueCase = [CaseInt 1] 596 | 597 | -- TODO: Stub implementation! 598 | -- Translating axioms seem to be problematic. For the other compiler they are 599 | -- defined in Agda.TypeChecking.Monad.Base. It is a field of 600 | -- `CompiledRepresentation`. We do not have this luxury. So what do we do? 601 | -- 602 | -- | Translates an axiom to a malfunction binding. Returns `Nothing` if the axiom is unmapped. 603 | compileAxiom :: 604 | MonadReader Env m => 605 | QName -- The name of the axiomm 606 | -> m (Maybe Binding) -- The resulting binding 607 | compileAxiom q = Just <$> namedBinding q x 608 | where 609 | x = fromMaybe unknownAxiom 610 | $ Map.lookup (show q') Primitive.axioms 611 | unknownAxiom = Mlambda [] $ errorT $ "Unknown axiom: " ++ show q' 612 | q' = last . qnameToList $ q 613 | 614 | -- | Translates a primitive to a malfunction binding. Returns `Nothing` if the primitive is unmapped. 615 | compilePrim 616 | :: MonadReader Env m => 617 | QName -- ^ The qname of the primitive 618 | -> String -- ^ The name of the primitive 619 | -> m (Maybe Binding) 620 | compilePrim q s = Just <$> namedBinding q x 621 | where 622 | x = fromMaybe unknownPrimitive 623 | $ Map.lookup s Primitive.primitives 624 | unknownPrimitive = Mlambda [] $ errorT $ "Unknown primitive: " ++ s 625 | 626 | namedBinding :: MonadReader Env m => QName -> Term -> m Binding 627 | namedBinding q t = (`Named`t) <$> nameToIdent q 628 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/Instances.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE DeriveDataTypeable #-} 2 | {-# LANGUAGE FlexibleInstances #-} 3 | {-# LANGUAGE StandaloneDeriving #-} 4 | {-# LANGUAGE TypeSynonymInstances #-} 5 | {-# OPTIONS_GHC -fno-warn-orphans #-} 6 | module Agda.Compiler.Malfunction.Instances where 7 | 8 | import Agda.Syntax.Common 9 | import qualified Agda.Syntax.Concrete as C 10 | import qualified Agda.Syntax.Fixity as F 11 | import Agda.Syntax.Literal 12 | import Agda.Syntax.Notation 13 | import Agda.Syntax.Position 14 | import Agda.Syntax.Treeless 15 | import Agda.Utils.FileName 16 | import Data.Data 17 | import Data.Graph 18 | 19 | deriving instance Data F.Fixity' 20 | deriving instance Data F.PrecedenceLevel 21 | deriving instance Data F.Associativity 22 | deriving instance Data MetaId 23 | deriving instance Data NameId 24 | deriving instance Data C.NamePart 25 | deriving instance Data C.Name 26 | deriving instance Data Hiding 27 | deriving instance Data Big 28 | deriving instance Data Relevance 29 | deriving instance Data Origin 30 | deriving instance Data ArgInfo 31 | deriving instance Data (Ranged RawName) 32 | deriving instance Data (Named_ Int) 33 | deriving instance Data (NamedArg Int) 34 | deriving instance Data GenPart 35 | deriving instance Data F.Fixity 36 | deriving instance Data AbsolutePath 37 | deriving instance Data (Position' ()) 38 | deriving instance Data (Interval' ()) 39 | deriving instance Data Range 40 | deriving instance Data TError 41 | deriving instance Data CaseType 42 | deriving instance Data ModuleName 43 | deriving instance Data Name 44 | deriving instance Data TAlt 45 | deriving instance Data Literal 46 | deriving instance Data TPrim 47 | deriving instance Data QName 48 | deriving instance Data TTerm 49 | 50 | deriving instance Typeable TTerm 51 | deriving instance Show vertex => Show (SCC vertex) 52 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/Main.hs: -------------------------------------------------------------------------------- 1 | module Agda.Compiler.Malfunction.Main (main) where 2 | 3 | import Agda.Main (runAgda) 4 | 5 | import qualified Agda.Compiler.Malfunction as Malfunction 6 | 7 | -- | Invokes the agda-compiler with the additional malfunction backend. 8 | main :: IO () 9 | main = runAgda [Malfunction.backend] 10 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/Primitive.hs: -------------------------------------------------------------------------------- 1 | {-# LANGUAGE OverloadedStrings #-} 2 | module Agda.Compiler.Malfunction.Primitive 3 | ( axioms 4 | , primitives 5 | ) where 6 | 7 | import Data.Map (Map) 8 | import qualified Data.Map as Map 9 | 10 | import Agda.Compiler.Malfunction.AST 11 | 12 | axioms :: Map String Term 13 | axioms = Map.fromList 14 | [ notMapped "Char.Char" 15 | , notMapped "IO.IO" 16 | , notMapped "String.String" 17 | , notMapped "IsOne" 18 | , notMapped "IsOne1" 19 | , notMapped "IsOne2" 20 | , notMapped "isOneEmpty" 21 | , notMapped "itIsOne" 22 | , notMapped "Level" 23 | , notMapped "Bytes" 24 | , notMapped "eqBytes" 25 | , notMapped "append" 26 | , notMapped "empty" 27 | , notMapped "erasedBottom" 28 | , notMapped "Unsafe._.trustme" 29 | , notMapped "exitWith'" 30 | , notMapped "getArgs" 31 | , notMapped "getChar" 32 | , notMapped "getProgName" 33 | , notMapped "ioReturn" 34 | , notMapped "putChar" 35 | , "putStrLn" |-> Mglobal ["Pervasives", "print_endline"] 36 | , "putStr" |-> Mglobal ["Pervasives", "print_string"] 37 | , notMapped "ioBind" 38 | , "primStringAppend" |-> primStringAppend 39 | ] 40 | where 41 | notMapped n = (n, Mlambda [] $ errorT $ "Axiom not yet mapped: " ++ n) 42 | 43 | primitives :: Map String Term 44 | primitives = Map.fromList 45 | -- Integer functions 46 | [ "primIntegerPlus" |-> binOp Add 47 | , "primIntegerMinus" |-> binOp Sub 48 | , "primIntegerTimes" |-> binOp Mul 49 | , "primIntegerDiv" |-> binOp Div 50 | , "primIntegerMod" |-> binOp Mod 51 | , "primIntegerEquality" |-> binOp Eq 52 | , "primIntegerLess" |-> binOp Lt 53 | , notMapped "primIntegerAbs" 54 | , notMapped "primNatToInteger" 55 | , "primShowInteger" |-> Mglobal ["Pervasives", "string_of_int"] 56 | 57 | -- Natural number functions 58 | , "primNatPlus" |-> binOp Add 59 | , "primNatMinus" |-> binOp Sub 60 | , "primNatTimes" |-> binOp Mul 61 | , "primNatDivSucAux" |-> binOp Div 62 | , "primNatModSucAux" |-> binOp Mod 63 | , "primNatEquality" |-> binOp Eq 64 | , "primNatLess" |-> binOp Lt 65 | 66 | -- Level functions 67 | , "primLevelZero" |-> zeroT 68 | , "primLevelSuc" |-> sucT 69 | , notMapped "primLevelMax" 70 | 71 | -- Floating point functions 72 | , notMapped "primNatToFloat" 73 | , notMapped "primFloatPlus" 74 | , notMapped "primFloatMinus" 75 | , notMapped "primFloatTimes" 76 | , notMapped "primFloatNegate" 77 | , notMapped "primFloatDiv" 78 | -- ASR (2016-09-29). We use bitwise equality for comparing Double 79 | -- because Haskell's Eq, which equates 0.0 and -0.0, allows to prove 80 | -- a contradiction (see Issue #2169). 81 | , notMapped "primFloatEquality" 82 | , notMapped "primFloatNumericalEquality" 83 | , notMapped "primFloatNumericalLess" 84 | , notMapped "primFloatSqrt" 85 | , notMapped "primRound" 86 | , notMapped "primFloor" 87 | , notMapped "primCeiling" 88 | , notMapped "primExp" 89 | , notMapped "primLog" 90 | , notMapped "primSin" 91 | , notMapped "primCos" 92 | , notMapped "primTan" 93 | , notMapped "primASin" 94 | , notMapped "primACos" 95 | , notMapped "primATan" 96 | , notMapped "primATan2" 97 | , notMapped "primShowFloat" 98 | 99 | -- Character functions 100 | , notMapped "primCharEquality" 101 | , notMapped "primIsLower" 102 | , notMapped "primIsDigit" 103 | , notMapped "primIsAlpha" 104 | , notMapped "primIsSpace" 105 | , notMapped "primIsAscii" 106 | , notMapped "primIsLatin1" 107 | , notMapped "primIsPrint" 108 | , notMapped "primIsHexDigit" 109 | , notMapped "primToUpper" 110 | , notMapped "primToLower" 111 | , notMapped "primCharToNat" 112 | , notMapped "primNatToChar" 113 | , notMapped "primShowChar" 114 | 115 | -- String functions 116 | , notMapped "primStringToList" 117 | , notMapped "primStringFromList" 118 | , "primStringAppend" |-> primStringAppend 119 | , notMapped "primStringEquality" 120 | , notMapped "primShowString" 121 | 122 | -- Other stuff 123 | , notMapped "primTrustMe" 124 | -- This needs to be force : A → ((x : A) → B x) → B x rather than seq because of call-by-name. 125 | , notMapped "primForce" 126 | , notMapped "primForceLemma" 127 | , notMapped "primQNameEquality" 128 | , notMapped "primQNameLess" 129 | , notMapped "primShowQName" 130 | , notMapped "primQNameFixity" 131 | , notMapped "primMetaEquality" 132 | , notMapped "primMetaLess" 133 | , notMapped "primShowMeta" 134 | ] 135 | where 136 | notMapped n = (n, Mlambda [] $ errorT $ "Primitive not yet mapped: " ++ n) 137 | 138 | (|->) :: a -> b -> (a, b) 139 | (|->) = (,) 140 | 141 | binOp :: BinaryIntOp -> Term 142 | binOp op = Mlambda ["a", "b"] (Mintop2 op TInt (Mvar "a") (Mvar "b")) 143 | 144 | zeroT :: Term 145 | zeroT = Mint (CInt 0) 146 | sucT :: Term 147 | sucT = Mlambda ["a"] (Mintop2 Add TInt (Mvar "a") (Mint (CInt 1))) 148 | 149 | -- FIXME: Copied from `Compiler` due to an otherwise cyclic dependency 150 | errorT :: String -> Term 151 | errorT err = Mapply (Mglobal ["Pervasives", "failwith"]) [Mstring err] 152 | 153 | primStringAppend :: Term 154 | primStringAppend = Mglobal ["Pervasives", "^"] 155 | -------------------------------------------------------------------------------- /src/full/Agda/Compiler/Malfunction/Run.hs: -------------------------------------------------------------------------------- 1 | module Agda.Compiler.Malfunction.Run 2 | ( compileRunPrint 3 | , compileModFile 4 | , compileRun 5 | , runModPrintInts 6 | , withPrintInts 7 | , runModFile 8 | , runMalfunction 9 | ) where 10 | 11 | import GHC.IO.Handle 12 | import System.FilePath 13 | import System.IO 14 | import System.IO.Temp 15 | import System.Process 16 | 17 | import Agda.Compiler.Malfunction.AST 18 | 19 | catMod :: Mod -> Handle -> IO () 20 | catMod m h = hPutStr h prog >> hFlush h 21 | where prog = prettyShow m 22 | 23 | -- | Compiles an existing file contain malfunction and writes the resulting 24 | -- executable to another file. 25 | compileModFile :: FilePath -> FilePath -> IO () 26 | compileModFile mlf exe = callProcess "malfunction" ["compile", mlf, "-o", exe] 27 | 28 | -- | Writes the module to a temporary files, compiles and executes it. 29 | runMod :: Mod -> IO String 30 | runMod t = withSystemTempFile "term.mlf" $ 31 | \tfp th -> do 32 | catMod t th 33 | runModFile' tfp th 34 | 35 | -- | Run .mlf 36 | runModFile :: FilePath -> IO String 37 | runModFile tfp = do 38 | th <- openFile tfp ReadMode 39 | runModFile' tfp th 40 | 41 | runModFile' :: FilePath -> Handle -> IO String 42 | runModFile' mlf th = 43 | withSystemTempFile "term.x" $ 44 | \xfp xh 45 | -> do 46 | compileModFile mlf xfp 47 | hClose th >> hClose xh 48 | readProcess xfp [] "" 49 | 50 | withPrintInts :: Mod -> [Ident] -> Mod 51 | withPrintInts (MMod bs expo) ids = MMod bs' expo 52 | where 53 | bs' = bs ++ map printInt ids 54 | printInt var = Unnamed $ Mapply (Mglobal ["Pervasives", "print_int"]) [Mvar var] 55 | 56 | runModPrintInts :: Mod -> [Ident] -> IO String 57 | runModPrintInts ids = runMod . withPrintInts ids 58 | 59 | -- | Compiles an to an executable and executes it. 60 | -- 61 | -- Note that this method uses the executable named `agda-ocaml` as registered with 62 | -- `stack`. 63 | compileRunPrint :: FilePath -> Ident -> IO String 64 | compileRunPrint agdap var = 65 | withSystemTempFile "module.mlf" $ 66 | \mlfp mlfh -> do 67 | callProcess "stack" ["exec", "agda-ocaml", "--", "-v0", "--mlf", agdap 68 | , "-o", mlfp, "--print-var", var] 69 | runModFile' mlfp mlfh 70 | 71 | compileRun :: FilePath -> IO String 72 | compileRun agdap = 73 | withSystemTempFile "module.mlf" $ 74 | \mlfp mlfh -> do 75 | callProcess "stack" ["exec", "agda-ocaml", "--", "-v0", "--mlf", agdap 76 | , "-o", mlfp] 77 | runModFile' mlfp mlfh 78 | 79 | 80 | -- FIXME: I do almost the same as existing functions in this module. 81 | -- 82 | -- I do realize that similar functionality exists in `Malfunction.Run` I didn't 83 | -- use this because it also prints some stuff to stdout, so I felt it was easier 84 | -- to just repeat these 3 lines of code. 85 | -- 86 | -- | `runMalfunction fp inp` calls the malfunction compiler on the input `inp` 87 | -- and places the output at `fp`. Assumes that the executable `malfunction` is 88 | -- in `PATH`. 89 | runMalfunction :: FilePath -> String -> IO () 90 | runMalfunction nm modl = takeBaseName nm `withSystemTempFile` \fp h -> do 91 | hPutStr h modl 92 | hFlush h 93 | callProcess "malfunction" ["compile", fp, "-o", nm] 94 | -------------------------------------------------------------------------------- /stack-8.0.2.yaml: -------------------------------------------------------------------------------- 1 | resolver: lts-8.5 2 | 3 | packages: 4 | - '.' 5 | - location: 6 | git: git@github.com:agda/agda.git 7 | commit: 8f19716ed22bade700ed5e182906175d16130207 8 | extra-dep: true 9 | - location: 10 | git: "https://github.com/lwm/tasty-discover.git" 11 | commit: "HEAD" 12 | extra-dep: true 13 | 14 | extra-deps: 15 | - tasty-th-0.1.4 16 | 17 | flags: {} 18 | 19 | extra-package-dbs: [] 20 | -------------------------------------------------------------------------------- /test/agda-ocaml/CompilerTest.hs: -------------------------------------------------------------------------------- 1 | module CompilerTest where 2 | 3 | -- TODO: Emacs keeps complaining that Test.Tasty.Discover is a member 4 | -- of a hidden package and keeps prompting me to add it to the .cabal-file. 5 | -- Solution M-x haskell-session-change-target -> agda-ocaml:test 6 | import qualified Agda.Syntax.Common as C 7 | import qualified Agda.Syntax.Concrete.Name as C 8 | import Agda.Syntax.Literal 9 | import Agda.Syntax.Treeless 10 | import qualified Data.Map as Map 11 | import GHC.Word 12 | import Test.Tasty 13 | import Test.Tasty.HUnit 14 | import Utils 15 | 16 | 17 | import Agda.Compiler.Malfunction.AST 18 | import Agda.Compiler.Malfunction.Compiler 19 | import Agda.Syntax.Common (NameId) 20 | 21 | 22 | translate'1 :: TTerm -> Term 23 | translate'1 = head . translateTerms' [] [] . pure 24 | 25 | translateTerms' :: [QName] -> [[(QName, Arity)]] -> [TTerm] -> [Term] 26 | translateTerms' ns qs = translateTerms (mkFakeEnv ns qs) 27 | 28 | translateDef' :: [QName] -> [[(QName, Arity)]] -> QName -> TTerm -> Binding 29 | translateDef' ns qs = translateDef (mkFakeEnv ns qs) 30 | 31 | mkFakeEnv :: [QName] -> [[(QName, Arity)]] -> Env 32 | mkFakeEnv ns cons = mkCompilerEnv2 ns cons -- ns (Map.unions (map mkMap cons)) 33 | 34 | mkMap :: [(QName, Arity)] -> Map.Map NameId ConRep 35 | mkMap qs = Map.fromList [ (qnameNameId qn, ConRep i arity) 36 | | (i, (qn, arity)) <- zip [0..] qs ] 37 | 38 | simpleName :: (C.NameId, String) -> Name 39 | simpleName (idf, concrete) = Name 40 | { nameId = idf 41 | , nameConcrete = C.Name undefined [C.Id concrete] 42 | , nameBindingSite = undefined 43 | , nameFixity = undefined 44 | } 45 | 46 | simpleQName :: [(C.NameId, String)] -> (C.NameId, String) -> QName 47 | simpleQName mods nm = QName { 48 | qnameModule = MName (map simpleName mods) 49 | , qnameName = simpleName nm 50 | } 51 | 52 | unitTests :: TestTree 53 | unitTests = testGroup "Compiler unit tests" test_translate 54 | 55 | -- TODO: Add this test-case: 56 | -- Agda: 57 | -- f = \ x -> g x 58 | -- g = f 59 | test_translate :: [TestTree] 60 | test_translate = 61 | -- Tests that the deBruijn index references the *closest* binding. 62 | [ testCase "sequenced lambda-expressions" 63 | $ translate'1 (TLam (TLam (TVar 0))) 64 | @?= Mlambda ["v0", "v1"] (Mvar "v1") 65 | , testCase "de Bruijn indices" $ 66 | translate'1 (TLam (TApp (TVar 0) [TLam (TVar 1)])) 67 | @?= Mlambda ["v0"] (Mapply (Mvar "v0") [Mlambda ["v1"] (Mvar "v0")]) 68 | , testCase "factorial" $ translateDef' [aName] [] aName (facTT aName) @?= facT anIden 69 | -- This test-case is a bit silly, since `TError TUnreachable` could be encoded 70 | -- as anything in malfunction. E.g. the function `f : ⊥ -> a` will never be 71 | -- applied to any arguments! 72 | , testCase "function from an uninhabited type" 73 | $ translate'1 (TError TUnreachable) 74 | @?= wildcardTerm 75 | -- , testCase "fst" 76 | -- $ translateDef' [[(aName, 1)]] aName (fstTT aName) @?= fstT aName 77 | -- This test-case fails for the same reason `fst` fails 78 | , testCase "non-nullary constructor application" 79 | $ runModExample constructorExample 80 | ] 81 | where 82 | (aName, anIden) = qnameAndIdent 0 "someId" 83 | 84 | facTT :: QName -> TTerm 85 | facTT qn = TLam (TCase 0 CTNat (TLet (TApp (TPrim PSub) 86 | [TVar 0,TLit (LitNat undefined 1)] 87 | ) (TApp (TPrim PMul) [TVar 1,TApp (TDef qn) [TVar 0]]) 88 | ) [TALit {aLit = LitNat undefined 0, aBody = TLit (LitNat undefined 1)}]) 89 | 90 | facT :: Ident -> Binding 91 | facT facIden = Recursive [(facIden, Mlambda ["v0"] 92 | (Mswitch (Mvar "v0") [([CaseInt 0],Mint (CBigint 1)) 93 | , ([CaseAnyInt,Deftag],Mlet [Named "v1" (Mintop2 Sub TInt 94 | (Mvar "v0") (Mint (CBigint 1))) 95 | ] (Mintop2 Mul TInt (Mvar "v0") 96 | (Mapply (Mvar facIden) [Mvar "v1"])))]))] 97 | 98 | -- fstT :: QName -> Binding 99 | -- fstT qn = Named (nameIdToIdent' (qnameNameId qn) "") (Mlambda ["v0"] (Mswitch (Mvar "v0") 100 | -- [([Tag 0],Mlet [Named "v1" (Mfield 0 (Mvar "v0")),Named "v2" (Mint (CBigint 666))] 101 | -- (Mvar "v1"))])) 102 | 103 | -- fstTT :: QName -> TTerm 104 | -- fstTT qn = TLam (TCase 0 (CTData qn) 105 | -- (TError TUnreachable) [ 106 | -- TACon {aCon = qn, aArity = 2, 107 | -- aBody = TVar 1}]) 108 | 109 | goldenTests :: TestTree 110 | goldenTests = testGroup "Compiler golden tests" 111 | [ mkGoldenGroup "FstSnd" ["a", "b"] 112 | , mkGoldenGroup "Factorial" ["a", "b"] 113 | , mkGoldenGroup "Constructor" ["one", "a"] 114 | , mkGoldenGroup "Index" ["l0", "l1", "l2"] 115 | , mkGoldenGroup "Insertion" ["l0", "l1"] 116 | , mkGoldenGroup "InsertionSort" ["l0", "l1", "l2"] 117 | , mkGoldenTest "RedBlack" (Left "short") 118 | ] 119 | 120 | runModExample :: ((Env, [(QName, TTerm)]), Mod) -> Assertion 121 | runModExample (inp, expected) = uncurry compile inp @?= expected 122 | 123 | constructorExample :: ((Env, [(QName, TTerm)]), Mod) 124 | constructorExample = (constructorTT qzero qsuc qf qone, constructorT if_ ione) 125 | where 126 | (qzero, izero) = qnameAndIdent 0 "zero" 127 | (qsuc, isuc) = qnameAndIdent 1 "suc" 128 | (qf, if_) = qnameAndIdent 2 "f" 129 | (qone, ione) = qnameAndIdent 3 "one" 130 | 131 | constructorTT :: QName -> QName -> QName -> QName -> (Env, [(QName, TTerm)]) 132 | constructorTT zero suc fQ oneQ = (env, bindings) 133 | where 134 | env = mkFakeEnv [zero, suc, fQ, oneQ] [[(zero, 0), (suc, 1)]] 135 | bindings = [(fQ, f), (oneQ, one)] 136 | f = TLam (TApp (TVar 0) [TCon zero]) 137 | one = TApp (TDef fQ) [TCon suc] 138 | 139 | constructorT :: Ident -> Ident -> Mod 140 | constructorT f one = MMod 141 | [ Named f (Mlambda ["v0"] (Mapply (Mvar "v0") [Mblock 0 []])) 142 | , Named one (Mapply (Mvar f) [Mlambda ["a"] (Mblock 1 [Mvar "a"])]) 143 | ] [] 144 | 145 | qnameAndIdent :: Word64 -> String -> (QName, Ident) 146 | qnameAndIdent nameid concrete = (simplerQName (nameid, concrete), 147 | nameIdToIdent' (C.NameId 0 nameid) (Just concrete)) 148 | where 149 | simplerQName :: (Word64, String) -> QName 150 | simplerQName (sndNameId, concrete) = simpleQName [] (fnameId, concrete) 151 | where fnameId = C.NameId 0 sndNameId 152 | 153 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Constructor.agda: -------------------------------------------------------------------------------- 1 | module Golden.Constructor where 2 | 3 | open import Agda.Builtin.Nat 4 | 5 | f : (Nat -> Nat) -> Nat 6 | f g = g zero 7 | 8 | fsuc = suc 9 | fzero = zero 10 | 11 | one = f suc 12 | a = fsuc fzero 13 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Constructor_a.golden: -------------------------------------------------------------------------------- 1 | 1 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Constructor_one.golden: -------------------------------------------------------------------------------- 1 | 1 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Factorial.agda: -------------------------------------------------------------------------------- 1 | module Golden.Factorial where 2 | 3 | open import Agda.Builtin.Nat 4 | 5 | fac : Nat -> Nat 6 | fac 0 = 1 7 | fac (suc n) = suc n * fac n 8 | 9 | a = fac 0 10 | b = fac 10 11 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Factorial_a.golden: -------------------------------------------------------------------------------- 1 | 1 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Factorial_b.golden: -------------------------------------------------------------------------------- 1 | 3628800 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/FstSnd.agda: -------------------------------------------------------------------------------- 1 | module Golden.FstSnd where 2 | 3 | open import Agda.Builtin.Nat 4 | 5 | data Tuple a b : Set where 6 | Tup : a -> b -> Tuple a b 7 | 8 | fst : {a b : Set} -> Tuple a b -> a 9 | fst (Tup x x₁) = x 10 | snd : {a b : Set} -> Tuple a b -> b 11 | snd (Tup x b) = b 12 | 13 | 14 | tup1 : Tuple Nat Nat 15 | tup1 = Tup 0 1 16 | 17 | a = fst tup1 18 | b = snd tup1 19 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/FstSnd_a.golden: -------------------------------------------------------------------------------- 1 | 0 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/FstSnd_b.golden: -------------------------------------------------------------------------------- 1 | 1 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Index.agda: -------------------------------------------------------------------------------- 1 | module Golden.Index where 2 | 3 | open import Agda.Builtin.Nat 4 | open import Agda.Builtin.List 5 | open import Agda.Builtin.Bool 6 | 7 | lst : List Nat 8 | lst = 1 ∷ 2 ∷ [] 9 | 10 | atDef : ∀ {a : Set} → a → List a -> Nat -> a 11 | atDef def (x ∷ l) zero = x 12 | atDef def (x ∷ l) (suc ix) = atDef def l ix 13 | atDef def _ _ = def 14 | 15 | l0 : Nat 16 | l0 = atDef 0 lst 0 17 | 18 | l1 : Nat 19 | l1 = atDef 0 lst 1 20 | 21 | l2 : Nat 22 | l2 = atDef 0 lst 2 23 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Index_l0.golden: -------------------------------------------------------------------------------- 1 | 1 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Index_l1.golden: -------------------------------------------------------------------------------- 1 | 2 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Index_l2.golden: -------------------------------------------------------------------------------- 1 | 0 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Insertion.agda: -------------------------------------------------------------------------------- 1 | module Golden.Insertion where 2 | 3 | open import Agda.Builtin.Nat 4 | open import Agda.Builtin.List 5 | open import Agda.Builtin.Bool 6 | 7 | insert : Nat -> List Nat -> List Nat 8 | insert a [] = a ∷ [] 9 | insert x (a ∷ b) with x < a 10 | ... | true = x ∷ a ∷ b 11 | ... | false = a ∷ (insert x b) 12 | 13 | lst : List Nat 14 | lst = 2 ∷ [] 15 | 16 | slst : List Nat 17 | slst = insert 3 lst 18 | 19 | atDef : ∀ {a : Set} → a → List a -> Nat -> a 20 | atDef def (x ∷ l) zero = x 21 | atDef def (x ∷ l) (suc ix) = atDef def l ix 22 | atDef def _ _ = def 23 | 24 | l0 : Nat 25 | l0 = atDef 0 slst 0 26 | 27 | l1 : Nat 28 | l1 = atDef 0 slst 1 29 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/InsertionSort.agda: -------------------------------------------------------------------------------- 1 | module Golden.InsertionSort where 2 | 3 | open import Agda.Builtin.Nat 4 | open import Agda.Builtin.List 5 | open import Agda.Builtin.Bool 6 | 7 | insert : Nat -> List Nat -> List Nat 8 | insert a [] = a ∷ [] 9 | insert x (a ∷ b) with x < a 10 | ... | true = x ∷ a ∷ b 11 | ... | false = a ∷ (insert x b) 12 | 13 | foldr : ∀ {a b : Set} → (a → b → b) → b → List a -> b 14 | foldr f ini [] = ini 15 | foldr f ini (x ∷ l) = f x (foldr f ini l) 16 | 17 | insertSort : List Nat -> List Nat 18 | insertSort = foldr insert [] 19 | 20 | atDef : ∀ {a : Set} → a → List a -> Nat -> a 21 | atDef def (x ∷ l) zero = x 22 | atDef def (x ∷ l) (suc ix) = atDef def l ix 23 | atDef def _ _ = def 24 | 25 | lst : List Nat 26 | lst = 4 ∷ 2 ∷ 7 ∷ [] 27 | 28 | slst : List Nat 29 | slst = insertSort lst 30 | 31 | l0 : Nat 32 | l0 = atDef 0 slst 0 33 | 34 | l1 : Nat 35 | l1 = atDef 0 slst 1 36 | 37 | l2 : Nat 38 | l2 = atDef 0 slst 2 39 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/InsertionSort_l0.golden: -------------------------------------------------------------------------------- 1 | 2 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/InsertionSort_l1.golden: -------------------------------------------------------------------------------- 1 | 4 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/InsertionSort_l2.golden: -------------------------------------------------------------------------------- 1 | 7 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Insertion_l0.golden: -------------------------------------------------------------------------------- 1 | 2 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/Insertion_l1.golden: -------------------------------------------------------------------------------- 1 | 3 -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/RedBlack.agda: -------------------------------------------------------------------------------- 1 | module Golden.RedBlack where 2 | 3 | open import Prelude hiding (insert) 4 | 5 | -- Version of comparison that lets us use instance search for the proof objects. 6 | data Comparison! {a} {A : Set a} (_<_ : A → A → Set a) (x y : A) : Set a where 7 | less : {{_ : x < y}} → Comparison! _<_ x y 8 | equal : {{_ : x ≡ y}} → Comparison! _<_ x y 9 | greater : {{_ : y < x}} → Comparison! _<_ x y 10 | 11 | compare! : ∀ {a} {A : Set a} {{_ : Ord A}} (x y : A) → Comparison! _<_ x y 12 | compare! x y = 13 | case compare x y of λ where 14 | (less lt) → less {{lt}} 15 | (equal eq) → equal {{eq}} 16 | (greater gt) → greater {{gt}} 17 | 18 | record Prf (A : Set) : Set where 19 | constructor ! 20 | field 21 | {{prf}} : A 22 | 23 | data Bound (A : Set) : Set where 24 | -∞ ∞ : Bound A 25 | # : A → Bound A 26 | 27 | module _ {A : Set} {{_ : Ord A}} where 28 | LessBound : Bound A → Bound A → Set 29 | LessBound ∞ _ = ⊥ 30 | LessBound _ ∞ = ⊤ 31 | LessBound _ -∞ = ⊥ 32 | LessBound -∞ _ = ⊤ 33 | LessBound (# x) (# y) = x < y 34 | 35 | instance 36 | OrdBound : Ord (Bound A) 37 | OrdBound = defaultOrd compareBound 38 | where 39 | compareBound : (a b : Bound A) → Comparison LessBound a b 40 | compareBound -∞ -∞ = equal refl 41 | compareBound -∞ ∞ = less _ 42 | compareBound -∞ (# x) = less _ 43 | compareBound ∞ -∞ = greater tt 44 | compareBound ∞ ∞ = equal refl 45 | compareBound ∞ (# x) = greater tt 46 | compareBound (# x) -∞ = greater tt 47 | compareBound (# x) ∞ = less tt 48 | compareBound (# x) (# y) with compare x y 49 | ... | less lt = less lt 50 | ... | greater gt = greater gt 51 | compareBound (# x) (# .x) | equal refl = equal refl 52 | 53 | Bounds : Set → Set 54 | Bounds A = Bound A × Bound A 55 | 56 | Rel : Set → Set₁ 57 | Rel A = Bounds A → Set 58 | 59 | module _ {A : Set} where 60 | PrfR : Rel A → Rel A 61 | PrfR R b = Prf (R b) 62 | 63 | data _∧_ (S T : Rel A) : Rel A where 64 | pivot : ∀ {l u} p → S (l , # p) → T (# p , u) → (S ∧ T) (l , u) 65 | 66 | pattern -⟨_⟩- p = pivot p ! ! 67 | pattern _⟨_⟩_ l p r = pivot p l r 68 | 69 | module _ {A : Set} {{_ : Ord A}} where 70 | Less : Rel A 71 | Less = PrfR (uncurry _<_) 72 | 73 | Bounded : Rel A 74 | Bounded = Less ∧ Less 75 | 76 | data Color : Set where 77 | red black : Color 78 | 79 | module _ (A : Set) {{_ : Ord A}} where 80 | 81 | data Tree′ (b : Bounds A) : Nat → Color → Set 82 | 83 | Tree : Nat → Color → Rel A 84 | Tree n c b = Tree′ b n c 85 | 86 | data Tree′ b where 87 | leaf′ : Less b → Tree 0 black b 88 | red : ∀ {n} → (Tree n black ∧ Tree n black) b → Tree n red b 89 | black : ∀ {c₁ c₂ n} → (Tree n c₁ ∧ Tree n c₂) b → Tree (suc n) black b 90 | 91 | pattern leaf = leaf′ ! 92 | 93 | data PreTree′ (b : Bounds A) (n : Nat) : Color → Set 94 | 95 | PreTree : Nat → Color → Bounds A → Set 96 | PreTree n c b = PreTree′ b n c 97 | 98 | data PreTree′ b n where 99 | red : ∀ {c₁ c₂} → (Tree n c₁ ∧ Tree n c₂) b → PreTree n red b 100 | maybe-black : ∀ {c} → Tree n c b → PreTree n black b 101 | 102 | pattern _b⟨_⟩_ l x r = black (l ⟨ x ⟩ r) 103 | pattern _r⟨_⟩_ l x r = red (l ⟨ x ⟩ r) 104 | pattern _pb⟨_⟩_ l x r = maybe-black (black (l ⟨ x ⟩ r)) 105 | pattern _pbr⟨_⟩_ l x r = maybe-black (red (l ⟨ x ⟩ r)) 106 | pattern _rbb⟨_⟩_ l x r = red {c₁ = black} {c₂ = black} (l ⟨ x ⟩ r) 107 | 108 | module _ {A : Set} {{_ : Ord A}} where 109 | 110 | balance-l : ∀ {b c₁ c₂ n} → 111 | (PreTree A n c₁ ∧ Tree A n c₂) b → 112 | PreTree A (suc n) black b 113 | balance-l (((l₁ r⟨ z ⟩ l₂) r⟨ x ⟩ l₃) ⟨ y ⟩ r) = 114 | (l₁ b⟨ z ⟩ l₂) pbr⟨ x ⟩ (l₃ b⟨ y ⟩ r) 115 | balance-l ((l₁ r⟨ z ⟩ (l₂ r⟨ x ⟩ l₃)) ⟨ y ⟩ r) = 116 | (l₁ b⟨ z ⟩ l₂) pbr⟨ x ⟩ (l₃ b⟨ y ⟩ r) 117 | balance-l ((l₁ rbb⟨ x ⟩ l₂) ⟨ y ⟩ r) = 118 | (l₁ r⟨ x ⟩ l₂) pb⟨ y ⟩ r 119 | balance-l (maybe-black l ⟨ y ⟩ r) = l pb⟨ y ⟩ r 120 | 121 | balance-r : ∀ {b c₁ c₂ n} → 122 | (Tree A n c₁ ∧ PreTree A n c₂) b → 123 | PreTree A (suc n) black b 124 | balance-r (l ⟨ y ⟩ ((r₁ r⟨ z ⟩ r₂) r⟨ x ⟩ r₃)) = 125 | (l b⟨ y ⟩ r₁) pbr⟨ z ⟩ (r₂ b⟨ x ⟩ r₃) 126 | balance-r (l ⟨ y ⟩ (r₁ r⟨ x ⟩ (r₂ r⟨ z ⟩ r₃))) = 127 | (l b⟨ y ⟩ r₁) pbr⟨ x ⟩ (r₂ b⟨ z ⟩ r₃) 128 | balance-r (l ⟨ y ⟩ (r₁ rbb⟨ x ⟩ r₂)) = 129 | l pb⟨ y ⟩ (r₁ r⟨ x ⟩ r₂) 130 | balance-r (l ⟨ y ⟩ maybe-black r) = l pb⟨ y ⟩ r 131 | 132 | ins : ∀ {b c n} → Bounded b → Tree A n c b → PreTree A n c b 133 | ins -⟨ x ⟩- leaf = leaf pbr⟨ x ⟩ leaf 134 | ins -⟨ x ⟩- (red (l ⟨ y ⟩ r)) = 135 | case compare! x y of λ where 136 | less → case ins -⟨ x ⟩- l of λ { (maybe-black l′) → l′ r⟨ y ⟩ r } 137 | greater → case ins -⟨ x ⟩- r of λ { (maybe-black r′) → l r⟨ y ⟩ r′ } 138 | equal → l r⟨ y ⟩ r 139 | ins -⟨ x ⟩- (black (l ⟨ y ⟩ r)) = 140 | case compare! x y of λ where 141 | less → balance-l (ins -⟨ x ⟩- l ⟨ y ⟩ r) 142 | greater → balance-r (l ⟨ y ⟩ ins -⟨ x ⟩- r) 143 | equal → l pb⟨ y ⟩ r 144 | 145 | data RedBlackTree (A : Set) {{_ : Ord A}} : Set where 146 | mkT : ∀ {n} → Tree A n black (-∞ , ∞) → RedBlackTree A 147 | 148 | module _ {A : Set} {{_ : Ord A}} where 149 | 150 | insert : A → RedBlackTree A → RedBlackTree A 151 | insert x (mkT t) with ins -⟨ x ⟩- t 152 | ... | l pbr⟨ y ⟩ r = mkT $ l b⟨ y ⟩ r 153 | ... | l pb⟨ y ⟩ r = mkT $ l b⟨ y ⟩ r 154 | ... | maybe-black leaf = mkT leaf 155 | 156 | fromList : List A → RedBlackTree A 157 | fromList = foldr insert (mkT leaf) 158 | 159 | toList′ : ∀ {b n c} → Tree′ A b n c → List A → List A 160 | toList′ (leaf′ _) xs = xs 161 | toList′ (l r⟨ x ⟩ r) xs = toList′ l (x ∷ toList′ r xs) 162 | toList′ (l b⟨ x ⟩ r) xs = toList′ l (x ∷ toList′ r xs) 163 | 164 | toList : RedBlackTree A → List A 165 | toList (mkT t) = toList′ t [] 166 | 167 | treeSort : List A → List A 168 | treeSort = toList ∘ fromList 169 | 170 | -- Example 171 | 172 | test : List Nat 173 | test = treeSort $ 5 ∷ 1 ∷ 2 ∷ 10 ∷ 13 ∷ 0 ∷ 141 ∷ 7 ∷ [] 174 | 175 | hest : String 176 | hest = show test 177 | 178 | main : IO Unit 179 | main = putStrLn hest 180 | -------------------------------------------------------------------------------- /test/agda-ocaml/Golden/RedBlack_short.golden: -------------------------------------------------------------------------------- 1 | [0, 1, 2, 5, 7, 10, 13, 141] 2 | -------------------------------------------------------------------------------- /test/agda-ocaml/Tasty.hs: -------------------------------------------------------------------------------- 1 | module Main (main) where 2 | 3 | import qualified CompilerTest as Compiler 4 | import Test.Tasty 5 | import System.Directory (withCurrentDirectory) 6 | 7 | -- Note that we need to change directory because of where the golden test-files are located 8 | main :: IO () 9 | main = withCurrentDirectory "test/agda-ocaml" $ defaultMain tests 10 | 11 | tests :: TestTree 12 | tests = testGroup "Tests" [unitTests, goldenTests] 13 | 14 | unitTests :: TestTree 15 | unitTests = testGroup "Unit tests" 16 | [Compiler.unitTests] 17 | 18 | goldenTests :: TestTree 19 | goldenTests = testGroup "Golden tests" 20 | [Compiler.goldenTests] 21 | -------------------------------------------------------------------------------- /test/agda-ocaml/Utils.hs: -------------------------------------------------------------------------------- 1 | module Utils where 2 | 3 | import qualified Data.ByteString.Lazy.Char8 as B 4 | import System.FilePath 5 | import Test.Tasty 6 | import Test.Tasty.Golden 7 | 8 | import Agda.Compiler.Malfunction.AST 9 | import Agda.Compiler.Malfunction.Run 10 | import Data.Either 11 | import Data.Either.Extra 12 | 13 | 14 | -- | .\/Golden\/FstSnd.agda .\/Golden\/FstSnd_a.golden 15 | -- mkdGoldenTest "FstSnd" "a" 16 | -- mkdGoldenTest "FstSnd" "b" 17 | mkGoldenTest :: String -> (Either String Ident) -> TestTree 18 | mkGoldenTest modPrefix testId = goldenVsString testName goldenPath result 19 | where 20 | testIdStr = fromEither testId 21 | result = case testId of 22 | Left str -> B.pack <$> compileRun agdap 23 | Right var -> B.pack <$> compileRunPrint agdap var' 24 | goldenPath = "Golden" modPrefix ++ "_" ++ testIdStr <.> "golden" 25 | var' = modPrefix ++ "." ++ testIdStr 26 | agdap = "Golden" modPrefix ++ ".agda" 27 | testName = modPrefix ++ " " ++ testIdStr 28 | 29 | mkGoldenGroup :: String -> [Ident] -> TestTree 30 | mkGoldenGroup modPrefix vars = testGroup modPrefix 31 | (map (mkGoldenTest modPrefix) (map Right vars)) 32 | --------------------------------------------------------------------------------