diff options
author | Nickolai Zeldovich <nickolai@csail.mit.edu> | 2015-05-23 12:22:48 -0700 |
---|---|---|
committer | Pierre Letouzey <pierre.letouzey@inria.fr> | 2015-06-22 14:15:00 +0200 |
commit | 5089872d20c9e3089676c9267a6394e99cca5121 (patch) | |
tree | 1e2df6141c807761304e25d3328c7f2ebf2ec5c8 /plugins/extraction | |
parent | 949d027ce8fa94b5c62f938b58c3f85d015b177b (diff) |
Add efficient extraction of [nat], [Z], and [string] to Haskell
This mirrors the existing extraction libraries for OCaml.
One wart: the extraction for [string] requires that the Haskell code
imports Data.Bits and Data.Char. Coq has no way to add extra import
statements to the extracted code. So we have to rely on the user to
somehow import these libraries (e.g., using the -pgmF ghc option).
See also https://coq.inria.fr/bugs/show_bug.cgi?id=4189
Message to github robot: this commit closes #65
Diffstat (limited to 'plugins/extraction')
-rw-r--r-- | plugins/extraction/ExtrHaskellNatInt.v | 13 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellNatInteger.v | 13 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellNatNum.v | 27 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellString.v | 38 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellZInt.v | 24 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellZInteger.v | 23 | ||||
-rw-r--r-- | plugins/extraction/ExtrHaskellZNum.v | 19 | ||||
-rw-r--r-- | plugins/extraction/vo.itarget | 9 |
8 files changed, 165 insertions, 1 deletions
diff --git a/plugins/extraction/ExtrHaskellNatInt.v b/plugins/extraction/ExtrHaskellNatInt.v new file mode 100644 index 000000000..e94e7d42b --- /dev/null +++ b/plugins/extraction/ExtrHaskellNatInt.v @@ -0,0 +1,13 @@ +(** Extraction of [nat] into Haskell's [Int] *) + +Require Import Arith. +Require Import ExtrHaskellNatNum. + +(** + * Disclaimer: trying to obtain efficient certified programs + * by extracting [nat] into [Int] is definitively *not* a good idea. + * See comments in [ExtrOcamlNatInt.v]. + *) + +Extract Inductive nat => "Prelude.Int" [ "0" "Prelude.succ" ] + "(\fO fS n -> if n Prelude.== 0 then fO () else fS (n Prelude.- 1))". diff --git a/plugins/extraction/ExtrHaskellNatInteger.v b/plugins/extraction/ExtrHaskellNatInteger.v new file mode 100644 index 000000000..038f0ed81 --- /dev/null +++ b/plugins/extraction/ExtrHaskellNatInteger.v @@ -0,0 +1,13 @@ +(** Extraction of [nat] into Haskell's [Integer] *) + +Require Import Arith. +Require Import ExtrHaskellNatNum. + +(** + * Disclaimer: trying to obtain efficient certified programs + * by extracting [nat] into [Integer] isn't necessarily a good idea. + * See comments in [ExtrOcamlNatInt.v]. +*) + +Extract Inductive nat => "Prelude.Integer" [ "0" "Prelude.succ" ] + "(\fO fS n -> if n Prelude.== 0 then fO () else fS (n Prelude.- 1))". diff --git a/plugins/extraction/ExtrHaskellNatNum.v b/plugins/extraction/ExtrHaskellNatNum.v new file mode 100644 index 000000000..979a1cdc4 --- /dev/null +++ b/plugins/extraction/ExtrHaskellNatNum.v @@ -0,0 +1,27 @@ +(** + * Efficient (but uncertified) extraction of usual [nat] functions + * into equivalent versions in Haskell's Prelude that are defined + * for any [Num] typeclass instances. Useful in combination with + * [Extract Inductive nat] that maps [nat] onto a Haskell type that + * implements [Num]. + *) + +Require Import Arith. +Require Import EqNat. + +Extract Inlined Constant Nat.add => "(Prelude.+)". +Extract Inlined Constant Nat.mul => "(Prelude.*)". +Extract Inlined Constant Nat.div => "Prelude.div". +Extract Inlined Constant Nat.max => "Prelude.max". +Extract Inlined Constant Nat.min => "Prelude.min". +Extract Inlined Constant Init.Nat.add => "(Prelude.+)". +Extract Inlined Constant Init.Nat.mul => "(Prelude.*)". +Extract Inlined Constant Compare_dec.lt_dec => "(Prelude.<)". +Extract Inlined Constant Compare_dec.leb => "(Prelude.<=)". +Extract Inlined Constant Compare_dec.le_lt_dec => "(Prelude.<=)". +Extract Inlined Constant EqNat.beq_nat => "(Prelude.==)". +Extract Inlined Constant EqNat.eq_nat_decide => "(Prelude.==)". +Extract Inlined Constant Peano_dec.eq_nat_dec => "(Prelude.==)". + +Extract Constant pred => "(\n -> Prelude.max 0 (Prelude.pred n))". +Extract Constant minus => "(\n m -> Prelude.max 0 (n Prelude.- m))". diff --git a/plugins/extraction/ExtrHaskellString.v b/plugins/extraction/ExtrHaskellString.v new file mode 100644 index 000000000..3558f4f25 --- /dev/null +++ b/plugins/extraction/ExtrHaskellString.v @@ -0,0 +1,38 @@ +(** + * Special handling of ascii and strings for extraction to Haskell. + *) + +Require Import Ascii. +Require Import String. + +(** + * At the moment, Coq's extraction has no way to add extra import + * statements to the extracted Haskell code. You will have to + * manually add: + * + * import qualified Data.Bits + * import qualified Data.Char + *) + +Extract Inductive ascii => "Prelude.Char" + [ "(\b0 b1 b2 b3 b4 b5 b6 b7 -> Data.Char.chr ( + (if b0 then Data.Bits.shiftL 1 0 else 0) Prelude.+ + (if b1 then Data.Bits.shiftL 1 1 else 0) Prelude.+ + (if b2 then Data.Bits.shiftL 1 2 else 0) Prelude.+ + (if b3 then Data.Bits.shiftL 1 3 else 0) Prelude.+ + (if b4 then Data.Bits.shiftL 1 4 else 0) Prelude.+ + (if b5 then Data.Bits.shiftL 1 5 else 0) Prelude.+ + (if b6 then Data.Bits.shiftL 1 6 else 0) Prelude.+ + (if b7 then Data.Bits.shiftL 1 7 else 0)))" ] + "(\f a -> f (Data.Bits.testBit (Data.Char.ord a) 0) + (Data.Bits.testBit (Data.Char.ord a) 1) + (Data.Bits.testBit (Data.Char.ord a) 2) + (Data.Bits.testBit (Data.Char.ord a) 3) + (Data.Bits.testBit (Data.Char.ord a) 4) + (Data.Bits.testBit (Data.Char.ord a) 5) + (Data.Bits.testBit (Data.Char.ord a) 6) + (Data.Bits.testBit (Data.Char.ord a) 7))". +Extract Inlined Constant Ascii.ascii_dec => "(Prelude.==)". + +Extract Inductive string => "Prelude.String" [ "([])" "(:)" ]. +Extract Inlined Constant String.string_dec => "(Prelude.==)". diff --git a/plugins/extraction/ExtrHaskellZInt.v b/plugins/extraction/ExtrHaskellZInt.v new file mode 100644 index 000000000..66690851a --- /dev/null +++ b/plugins/extraction/ExtrHaskellZInt.v @@ -0,0 +1,24 @@ +(** Extraction of [Z] into Haskell's [Int] *) + +Require Import ZArith. +Require Import ExtrHaskellZNum. + +(** + * Disclaimer: trying to obtain efficient certified programs + * by extracting [Z] into [Int] is definitively *not* a good idea. + * See comments in [ExtrOcamlNatInt.v]. + *) + +Extract Inductive positive => "Prelude.Int" [ + "(\x -> 2 Prelude.* x Prelude.+ 1)" + "(\x -> 2 Prelude.* x)" + "1" ] + "(\fI fO fH n -> if n Prelude.== 1 then fH () else + if Prelude.odd n + then fI (n `Prelude.div` 2) + else fO (n `Prelude.div` 2))". + +Extract Inductive Z => "Prelude.Int" [ "0" "(\x -> x)" "Prelude.negate" ] + "(\fO fP fN n -> if n Prelude.== 0 then fO () else + if n Prelude.> 0 then fP n else + fN (Prelude.negate n))". diff --git a/plugins/extraction/ExtrHaskellZInteger.v b/plugins/extraction/ExtrHaskellZInteger.v new file mode 100644 index 000000000..f192f16ee --- /dev/null +++ b/plugins/extraction/ExtrHaskellZInteger.v @@ -0,0 +1,23 @@ +(** Extraction of [Z] into Haskell's [Integer] *) + +Require Import ZArith. +Require Import ExtrHaskellZNum. + +(** Disclaimer: trying to obtain efficient certified programs + by extracting [Z] into [Integer] isn't necessarily a good idea. + See comments in [ExtrOcamlNatInt.v]. +*) + +Extract Inductive positive => "Prelude.Integer" [ + "(\x -> 2 Prelude.* x Prelude.+ 1)" + "(\x -> 2 Prelude.* x)" + "1" ] + "(\fI fO fH n -> if n Prelude.== 1 then fH () else + if Prelude.odd n + then fI (n `Prelude.div` 2) + else fO (n `Prelude.div` 2))". + +Extract Inductive Z => "Prelude.Integer" [ "0" "(\x -> x)" "Prelude.negate" ] + "(\fO fP fN n -> if n Prelude.== 0 then fO () else + if n Prelude.> 0 then fP n else + fN (Prelude.negate n))". diff --git a/plugins/extraction/ExtrHaskellZNum.v b/plugins/extraction/ExtrHaskellZNum.v new file mode 100644 index 000000000..3f645db9b --- /dev/null +++ b/plugins/extraction/ExtrHaskellZNum.v @@ -0,0 +1,19 @@ +(** + * Efficient (but uncertified) extraction of usual [Z] functions + * into equivalent versions in Haskell's Prelude that are defined + * for any [Num] typeclass instances. Useful in combination with + * [Extract Inductive Z] that maps [Z] onto a Haskell type that + * implements [Num]. + *) + +Require Import ZArith. +Require Import EqNat. + +Extract Inlined Constant Z.add => "(Prelude.+)". +Extract Inlined Constant Z.sub => "(Prelude.-)". +Extract Inlined Constant Z.mul => "(Prelude.*)". +Extract Inlined Constant Z.div => "Prelude.div". +Extract Inlined Constant Z.max => "Prelude.max". +Extract Inlined Constant Z.min => "Prelude.min". +Extract Inlined Constant Z_ge_lt_dec => "(Prelude.>=)". +Extract Inlined Constant Z_gt_le_dec => "(Prelude.>)". diff --git a/plugins/extraction/vo.itarget b/plugins/extraction/vo.itarget index f04890480..9c30c5eb3 100644 --- a/plugins/extraction/vo.itarget +++ b/plugins/extraction/vo.itarget @@ -1,4 +1,11 @@ ExtrHaskellBasic.vo +ExtrHaskellNatNum.vo +ExtrHaskellNatInt.vo +ExtrHaskellNatInteger.vo +ExtrHaskellZNum.vo +ExtrHaskellZInt.vo +ExtrHaskellZInteger.vo +ExtrHaskellString.vo ExtrOcamlBasic.vo ExtrOcamlIntConv.vo ExtrOcamlBigIntConv.vo @@ -6,4 +13,4 @@ ExtrOcamlNatInt.vo ExtrOcamlNatBigInt.vo ExtrOcamlZInt.vo ExtrOcamlZBigInt.vo -ExtrOcamlString.vo
\ No newline at end of file +ExtrOcamlString.vo |