blob: 66690851a7ece2d4ea94f44b1ab5e24b6f2bd847 (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
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))".
|