summaryrefslogtreecommitdiff
path: root/plugins/extraction/ExtrHaskellNatInt.v
blob: e94e7d42bd912d8a12e42e598a55b8a258a85e76 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
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))".